@@ -58,6 +58,7 @@ module _ {_∼_ : Rel A ℓ} where
58
58
module _ (_∼_ : Rel A ℓ) where
59
59
private
60
60
_∼⁺_ = TransClosure _∼_
61
+ module ∼⊆∼⁺ = Subrelation {_<₂_ = _∼⁺_} [_]
61
62
62
63
reflexive : Reflexive _∼_ → Reflexive _∼⁺_
63
64
reflexive refl = [ refl ]
@@ -69,17 +70,21 @@ module _ (_∼_ : Rel A ℓ) where
69
70
transitive : Transitive _∼⁺_
70
71
transitive = _++_
71
72
72
- wellFounded : WellFounded _∼_ → WellFounded _∼⁺_
73
- wellFounded wf = λ x → acc (accessible′ (wf x))
74
- where
75
- downwardsClosed : ∀ {x y} → Acc _∼⁺_ y → x ∼ y → Acc _∼⁺_ x
76
- downwardsClosed (acc rec) x∼y = acc (λ z z∼x → rec z (z∼x ∷ʳ x∼y))
73
+ accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x
74
+ accessible⁻ = ∼⊆∼⁺.accessible
77
75
78
- accessible′ : ∀ {x} → Acc _∼_ x → WfRec _∼⁺_ (Acc _∼⁺_) x
79
- accessible′ (acc rec) y [ y∼x ] = acc (accessible′ (rec y y∼x))
80
- accessible′ acc[x] y (y∼z ∷ z∼⁺x) =
81
- downwardsClosed (accessible′ acc[x] _ z∼⁺x) y∼z
76
+ wellFounded⁻ : WellFounded _∼⁺_ → WellFounded _∼_
77
+ wellFounded⁻ = ∼⊆∼⁺.wellFounded
82
78
79
+ accessible : ∀ {x} → Acc _∼_ x → Acc _∼⁺_ x
80
+ accessible acc[x] = acc (wf-acc acc[x])
81
+ where
82
+ wf-acc : ∀ {x} → Acc _∼_ x → WfRec _∼⁺_ (Acc _∼⁺_) x
83
+ wf-acc (acc rec) _ [ y∼x ] = acc (wf-acc (rec _ y∼x))
84
+ wf-acc acc[x] _ (y∼z ∷ z∼⁺x) = acc-inverse (wf-acc acc[x] _ z∼⁺x) _ [ y∼z ]
85
+
86
+ wellFounded : WellFounded _∼_ → WellFounded _∼⁺_
87
+ wellFounded wf x = accessible (wf x)
83
88
84
89
85
90
------------------------------------------------------------------------
0 commit comments