Skip to content

Commit aff1a42

Browse files
authored
WellFounded proofs for transitive closure (#2082)
1 parent fceac19 commit aff1a42

File tree

2 files changed

+21
-9
lines changed

2 files changed

+21
-9
lines changed

CHANGELOG.md

Lines changed: 7 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -2951,6 +2951,13 @@ Other minor changes
29512951
f Preserves₂ ≈₁ ⟶ ≈₁ ⟶ ≈₂
29522952
```
29532953

2954+
* Added new proofs to `Relation.Binary.Construct.Closure.Transitive`:
2955+
```
2956+
accessible⁻ : ∀ {x} → Acc _∼⁺_ x → Acc _∼_ x
2957+
wellFounded⁻ : WellFounded _∼⁺_ → WellFounded _∼_
2958+
accessible : ∀ {x} → Acc _∼_ x → Acc _∼⁺_ x
2959+
```
2960+
29542961
* Added new operations in `Relation.Binary.PropositionalEquality.Properties`:
29552962
```
29562963
J : (B : (y : A) → x ≡ y → Set b) (p : x ≡ y) → B x refl → B y p

src/Relation/Binary/Construct/Closure/Transitive.agda

Lines changed: 14 additions & 9 deletions
Original file line numberDiff line numberDiff line change
@@ -58,6 +58,7 @@ module _ {_∼_ : Rel A ℓ} where
5858
module _ (_∼_ : Rel A ℓ) where
5959
private
6060
_∼⁺_ = TransClosure _∼_
61+
module ∼⊆∼⁺ = Subrelation {_<₂_ = _∼⁺_} [_]
6162

6263
reflexive : Reflexive _∼_ Reflexive _∼⁺_
6364
reflexive refl = [ refl ]
@@ -69,17 +70,21 @@ module _ (_∼_ : Rel A ℓ) where
6970
transitive : Transitive _∼⁺_
7071
transitive = _++_
7172

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
7775

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
8278

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)
8388

8489

8590
------------------------------------------------------------------------

0 commit comments

Comments
 (0)