Skip to content

Commit 08a89c7

Browse files
jamesmckinnaandreasabel
authored andcommitted
refactored proofs from #2023 (#2301)
1 parent 6c925fa commit 08a89c7

File tree

3 files changed

+27
-18
lines changed

3 files changed

+27
-18
lines changed

CHANGELOG.md

Lines changed: 20 additions & 15 deletions
Original file line numberDiff line numberDiff line change
@@ -141,28 +141,28 @@ Additions to existing modules
141141

142142
* In `Data.List.Relation.Ternary.Appending.Setoid.Properties`:
143143
```agda
144-
through→ : ∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs →
144+
through→ : ∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs →
145145
∃[ ys ] Appending as bs ys × Pointwise _≈_ ys cs
146-
through← : ∃[ ys ] Appending as bs ys × Pointwise _≈_ ys cs →
146+
through← : ∃[ ys ] Appending as bs ys × Pointwise _≈_ ys cs →
147147
∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs
148-
assoc→ : ∃[ xs ] Appending as bs xs × Appending xs cs ds →
148+
assoc→ : ∃[ xs ] Appending as bs xs × Appending xs cs ds →
149149
∃[ ys ] Appending bs cs ys × Appending as ys ds
150150
```
151151

152152
* In `Data.List.Relation.Ternary.Appending.Properties`:
153153
```agda
154-
through→ : (R ⇒ (S ; T)) → ((U ; V) ⇒ (W ; T)) →
155-
∃[ xs ] Pointwise U as xs × Appending V R xs bs cs →
156-
∃[ ys ] Appending W S as bs ys × Pointwise T ys cs
157-
through← : ((R ; S) ⇒ T) → ((U ; S) ⇒ (V ; W)) →
158-
∃[ ys ] Appending U R as bs ys × Pointwise S ys cs →
159-
∃[ xs ] Pointwise V as xs × Appending W T xs bs cs
160-
assoc→ : (R ⇒ (S ; T)) → ((U ; V) ⇒ (W ; T)) → ((Y ; V) ⇒ X) →
161-
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds →
162-
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds
163-
assoc← : ((S ; T) ⇒ R) → ((W ; T) ⇒ (U ; V)) → (X ⇒ (Y ; V)) →
164-
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds →
165-
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds
154+
through→ : (R ⇒ (S ; T)) → ((U ; V) ⇒ (W ; T)) →
155+
∃[ xs ] Pointwise U as xs × Appending V R xs bs cs →
156+
∃[ ys ] Appending W S as bs ys × Pointwise T ys cs
157+
through← : ((R ; S) ⇒ T) → ((U ; S) ⇒ (V ; W)) →
158+
∃[ ys ] Appending U R as bs ys × Pointwise S ys cs →
159+
∃[ xs ] Pointwise V as xs × Appending W T xs bs cs
160+
assoc→ : (R ⇒ (S ; T)) → ((U ; V) ⇒ (W ; T)) → ((Y ; V) ⇒ X) →
161+
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds →
162+
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds
163+
assoc← : ((S ; T) ⇒ R) → ((W ; T) ⇒ (U ; V)) → (X ⇒ (Y ; V)) →
164+
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds →
165+
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds
166166
```
167167

168168
* In `Data.List.Relation.Binary.Pointwise.Base`:
@@ -210,6 +210,11 @@ Additions to existing modules
210210
* In `Function.Bundles`, added `_⟶ₛ_` as a synonym for `Func` that can
211211
be used infix.
212212

213+
* Added new proofs in `Relation.Binary.Construct.Composition`:
214+
```agda
215+
transitive⇒≈;≈⊆≈ : Transitive ≈ → (≈ ; ≈) ⇒ ≈
216+
```
217+
213218
* Added new definitions in `Relation.Binary.Definitions`
214219
```
215220
Stable _∼_ = ∀ x y → Nullary.Stable (x ∼ y)

src/Relation/Binary/Construct/Composition.agda

Lines changed: 3 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -82,3 +82,6 @@ module _ (L : Rel A ℓ₁) (R : Rel A ℓ₂) (comm : R ; L ⇒ L ; R) where
8282
; trans = transitive Oˡ.trans Oʳ.trans
8383
}
8484
where module = IsPreorder Oˡ; module = IsPreorder Oʳ
85+
86+
transitive⇒≈;≈⊆≈ : (≈ : Rel A ℓ) Transitive ≈ (≈ ; ≈) ⇒ ≈
87+
transitive⇒≈;≈⊆≈ _ trans (_ , l , r) = trans l r

src/Relation/Binary/Properties/Setoid.agda

Lines changed: 4 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -15,7 +15,8 @@ open import Relation.Binary.Bundles using (Setoid; Preorder; Poset)
1515
open import Relation.Binary.Definitions
1616
using (Symmetric; _Respectsˡ_; _Respectsʳ_; _Respects₂_)
1717
open import Relation.Binary.Structures using (IsPreorder; IsPartialOrder)
18-
open import Relation.Binary.Construct.Composition using (_;_)
18+
open import Relation.Binary.Construct.Composition
19+
using (_;_; impliesˡ; transitive⇒≈;≈⊆≈)
1920

2021
module Relation.Binary.Properties.Setoid {a ℓ} (S : Setoid a ℓ) where
2122

@@ -83,10 +84,10 @@ preorder = record
8384
-- Equality is closed under composition
8485

8586
≈;≈⇒≈ : _≈_ ; _≈_ ⇒ _≈_
86-
≈;≈⇒≈ (_ , p , q) = trans p q
87+
≈;≈⇒≈ = transitive⇒≈;≈⊆≈ _ trans
8788

8889
≈⇒≈;≈ : _≈_ ⇒ _≈_ ; _≈_
89-
≈⇒≈;≈ q = _ , q , refl
90+
≈⇒≈;≈ = impliesˡ _≈_ _≈_ refl id
9091

9192
------------------------------------------------------------------------
9293
-- Other properties

0 commit comments

Comments
 (0)