Skip to content

Commit 67bed00

Browse files
[ new ] associativity of Appending (#2023)
* [ new ] associativity of Appending * Removed unneeded variable * Renamed Product module --------- Co-authored-by: MatthewDaggitt <[email protected]>
1 parent 2fe12da commit 67bed00

File tree

5 files changed

+157
-39
lines changed

5 files changed

+157
-39
lines changed

CHANGELOG.md

Lines changed: 40 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -13,6 +13,9 @@ Bug-fixes
1313
was mistakenly applied to the level of the type `A` instead of the
1414
variable `x` of type `A`.
1515

16+
* Module `Data.List.Relation.Ternary.Appending.Setoid.Properties` no longer
17+
exports the `Setoid` module under the alias `S`.
18+
1619
Non-backwards compatible changes
1720
--------------------------------
1821

@@ -131,6 +134,37 @@ Additions to existing modules
131134
tabulate⁺-< : (i < j → R (f i) (f j)) → AllPairs R (tabulate f)
132135
```
133136

137+
* In `Data.List.Relation.Ternary.Appending.Setoid.Properties`:
138+
```agda
139+
through→ : ∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs →
140+
∃[ ys ] Appending as bs ys × Pointwise _≈_ ys cs
141+
through← : ∃[ ys ] Appending as bs ys × Pointwise _≈_ ys cs →
142+
∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs
143+
assoc→ : ∃[ xs ] Appending as bs xs × Appending xs cs ds →
144+
∃[ ys ] Appending bs cs ys × Appending as ys ds
145+
```
146+
147+
* In `Data.List.Relation.Ternary.Appending.Properties`:
148+
```agda
149+
through→ : (R ⇒ (S ; T)) → ((U ; V) ⇒ (W ; T)) →
150+
∃[ xs ] Pointwise U as xs × Appending V R xs bs cs →
151+
∃[ ys ] Appending W S as bs ys × Pointwise T ys cs
152+
through← : ((R ; S) ⇒ T) → ((U ; S) ⇒ (V ; W)) →
153+
∃[ ys ] Appending U R as bs ys × Pointwise S ys cs →
154+
∃[ xs ] Pointwise V as xs × Appending W T xs bs cs
155+
assoc→ : (R ⇒ (S ; T)) → ((U ; V) ⇒ (W ; T)) → ((Y ; V) ⇒ X) →
156+
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds →
157+
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds
158+
assoc← : ((S ; T) ⇒ R) → ((W ; T) ⇒ (U ; V)) → (X ⇒ (Y ; V)) →
159+
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds →
160+
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds
161+
```
162+
163+
* In `Data.List.Relation.Binary.Pointwise.Base`:
164+
```agda
165+
unzip : Pointwise (R ; S) ⇒ (Pointwise R ; Pointwise S)
166+
```
167+
134168
* In `Data.Maybe.Relation.Binary.Pointwise`:
135169
```agda
136170
pointwise⊆any : Pointwise R (just x) ⊆ Any (R x)
@@ -176,6 +210,12 @@ Additions to existing modules
176210
Stable : Pred A ℓ → Set _
177211
```
178212

213+
* Added new proofs in `Relation.Binary.Properties.Setoid`:
214+
```agda
215+
≈;≈⇒≈ : _≈_ ; _≈_ ⇒ _≈_
216+
≈⇒≈;≈ : _≈_ ⇒ _≈_ ; _≈_
217+
```
218+
179219
* Added new definitions in `Relation.Nullary`
180220
```
181221
Recomputable : Set _

src/Data/List/Relation/Binary/Pointwise/Base.agda

Lines changed: 8 additions & 3 deletions
Original file line numberDiff line numberDiff line change
@@ -8,16 +8,16 @@
88

99
module Data.List.Relation.Binary.Pointwise.Base where
1010

11-
open import Data.Product.Base using (_×_; <_,_>)
11+
open import Data.Product.Base as Product using (_×_; _,_; <_,_>; ∃-syntax)
1212
open import Data.List.Base using (List; []; _∷_)
1313
open import Level using (Level; _⊔_)
1414
open import Relation.Binary.Core using (REL; _⇒_)
15+
open import Relation.Binary.Construct.Composition using (_;_)
1516

1617
private
1718
variable
1819
a b c ℓ : Level
19-
A : Set a
20-
B : Set b
20+
A B : Set a
2121
x y : A
2222
xs ys : List A
2323
R S : REL A B ℓ
@@ -58,3 +58,8 @@ rec P c n (Rxy ∷ Rxsys) = c Rxy (rec P c n Rxsys)
5858
map : R ⇒ S Pointwise R ⇒ Pointwise S
5959
map R⇒S [] = []
6060
map R⇒S (Rxy ∷ Rxsys) = R⇒S Rxy ∷ map R⇒S Rxsys
61+
62+
unzip : Pointwise (R ; S) ⇒ (Pointwise R ; Pointwise S)
63+
unzip [] = [] , [] , []
64+
unzip ((y , r , s) ∷ xs∼ys) =
65+
Product.map (y ∷_) (Product.map (r ∷_) (s ∷_)) (unzip xs∼ys)

src/Data/List/Relation/Ternary/Appending/Properties.agda

Lines changed: 70 additions & 29 deletions
Original file line numberDiff line numberDiff line change
@@ -10,48 +10,89 @@ module Data.List.Relation.Ternary.Appending.Properties where
1010

1111
open import Data.List.Base using (List; [])
1212
open import Data.List.Relation.Ternary.Appending
13+
open import Data.List.Relation.Binary.Pointwise as Pw using (Pointwise; []; _∷_)
14+
open import Data.Product.Base as Product using (∃-syntax; _×_; _,_)
15+
open import Function.Base using (id)
1316
open import Data.List.Relation.Binary.Pointwise.Base as Pw using (Pointwise; []; _∷_)
1417
open import Data.List.Relation.Binary.Pointwise.Properties as Pw using (transitive)
1518
open import Level using (Level)
16-
open import Relation.Binary.Core using (REL; Rel)
19+
open import Relation.Binary.Core using (REL; Rel; _⇒_)
1720
open import Relation.Binary.Definitions using (Trans)
1821
open import Relation.Binary.PropositionalEquality.Core using (_≡_; refl)
22+
open import Relation.Binary.Construct.Composition using (_;_)
1923

2024
private
2125
variable
22-
a a′ b b′ c c′ l r : Level
23-
A : Set a
24-
A′ : Set a′
25-
B : Set b
26-
B′ : Set b′
27-
C : Set c
28-
C′ : Set c′
29-
L : REL A C l
30-
R : REL B C r
31-
as : List A
32-
bs : List B
33-
cs : List C
34-
35-
module _ {e} {E : REL C C′ e} {L′ : REL A C′ l} {R′ : REL B C′ r}
36-
(LEL′ : Trans L E L′) (RER′ : Trans R E R′)
37-
where
26+
a ℓ l r : Level
27+
A A′ B B′ C C′ D D′ : Set a
28+
R S T U V W X Y : REL A B ℓ
29+
as bs cs ds : List A
30+
31+
module _ (RST : Trans R S T) (USV : Trans U S V) where
3832

39-
respʳ-≋ : {cs′} Appending L R as bs cs Pointwise E cs cs′ Appending L′ R′ as bs cs′
40-
respʳ-≋ ([]++ rs) es = []++ Pw.transitive RER′ rs es
41-
respʳ-≋ (l ∷ lrs) (e ∷ es) = LEL′ l e ∷ respʳ-≋ lrs es
33+
respʳ-≋ : Appending R U as bs cs Pointwise S cs ds Appending T V as bs ds
34+
respʳ-≋ ([]++ rs) es = []++ Pw.transitive USV rs es
35+
respʳ-≋ (l ∷ lrs) (e ∷ es) = RST l e ∷ respʳ-≋ lrs es
4236

43-
module _ {eᴬ eᴮ} {Eᴬ : REL A′ A eᴬ} {Eᴮ : REL B′ B eᴮ}
44-
{L′ : REL A′ C l} (ELL′ : Trans Eᴬ L L′)
45-
{R′ : REL B′ C r} (ERR′ : Trans Eᴮ R R′)
37+
module _ {T : REL A B l} (RST : Trans R S T)
38+
{W : REL A B r} (ERW : Trans U V W)
4639
where
4740

48-
respˡ-≋ : {as′ bs′} Pointwise Eᴬ as′ as Pointwise Eᴮ bs′ bs
49-
Appending L R as bs cs Appending L′ R′ as′ bs′ cs
50-
respˡ-≋ [] esʳ ([]++ rs) = []++ Pw.transitive ERR′ esʳ rs
51-
respˡ-≋ (eˡ ∷ esˡ) esʳ (l ∷ lrs) = ELL′ eˡ l ∷ respˡ-≋ esˡ esʳ lrs
41+
respˡ-≋ : {as′ bs′} Pointwise R as′ as Pointwise U bs′ bs
42+
Appending S V as bs cs Appending T W as′ bs′ cs
43+
respˡ-≋ [] esʳ ([]++ rs) = []++ Pw.transitive ERW esʳ rs
44+
respˡ-≋ (eˡ ∷ esˡ) esʳ (l ∷ lrs) = RST eˡ l ∷ respˡ-≋ esˡ esʳ lrs
5245

53-
conicalˡ : Appending L R as bs [] as ≡ []
46+
conicalˡ : Appending R S as bs [] as ≡ []
5447
conicalˡ ([]++ rs) = refl
5548

56-
conicalʳ : Appending L R as bs [] bs ≡ []
49+
conicalʳ : Appending R S as bs [] bs ≡ []
5750
conicalʳ ([]++ []) = refl
51+
52+
through→ :
53+
(R ⇒ (S ; T))
54+
((U ; V) ⇒ (W ; T))
55+
∃[ xs ] Pointwise U as xs × Appending V R xs bs cs
56+
∃[ ys ] Appending W S as bs ys × Pointwise T ys cs
57+
through→ f g (_ , [] , []++ rs) =
58+
let _ , rs′ , ps′ = Pw.unzip (Pw.map f rs) in
59+
_ , []++ rs′ , ps′
60+
through→ f g (_ , p ∷ ps , l ∷ lrs) =
61+
let _ , l′ , p′ = g (_ , p , l) in
62+
Product.map _ (Product.map (l′ ∷_) (p′ ∷_)) (through→ f g (_ , ps , lrs))
63+
64+
through← :
65+
((R ; S) ⇒ T)
66+
((U ; S) ⇒ (V ; W))
67+
∃[ ys ] Appending U R as bs ys × Pointwise S ys cs
68+
∃[ xs ] Pointwise V as xs × Appending W T xs bs cs
69+
through← f g (_ , []++ rs′ , ps′) =
70+
_ , [] , []++ (Pw.transitive (λ r′ p′ f (_ , r′ , p′)) rs′ ps′)
71+
through← f g (_ , l′ ∷ lrs′ , p′ ∷ ps′) =
72+
let _ , p , l = g (_ , l′ , p′) in
73+
Product.map _ (Product.map (p ∷_) (l ∷_)) (through← f g (_ , lrs′ , ps′))
74+
75+
assoc→ :
76+
(R ⇒ (S ; T))
77+
((U ; V) ⇒ (W ; T))
78+
((Y ; V) ⇒ X)
79+
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds
80+
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds
81+
assoc→ f g h (_ , []++ rs , lrs′) =
82+
let _ , mss , ss′ = through→ f g (_ , rs , lrs′) in
83+
_ , mss , []++ ss′
84+
assoc→ f g h (_ , l ∷ lrs , l′ ∷ lrs′) =
85+
Product.map₂ (Product.map₂ (h (_ , l , l′) ∷_)) (assoc→ f g h (_ , lrs , lrs′))
86+
87+
assoc← :
88+
((S ; T) ⇒ R)
89+
((W ; T) ⇒ (U ; V))
90+
(X ⇒ (Y ; V))
91+
∃[ ys ] Appending W S bs cs ys × Appending X T as ys ds
92+
∃[ xs ] Appending Y U as bs xs × Appending V R xs cs ds
93+
assoc← f g h (_ , mss , []++ ss′) =
94+
let _ , rs , lrs′ = through← f g (_ , mss , ss′) in
95+
_ , []++ rs , lrs′
96+
assoc← f g h (_ , mss , m′ ∷ mss′) =
97+
let _ , l , l′ = h m′ in
98+
Product.map _ (Product.map (l ∷_) (l′ ∷_)) (assoc← f g h (_ , mss , mss′))

src/Data/List/Relation/Ternary/Appending/Setoid/Properties.agda

Lines changed: 28 additions & 7 deletions
Original file line numberDiff line numberDiff line change
@@ -8,27 +8,33 @@
88

99
open import Relation.Binary.Bundles using (Setoid)
1010

11-
module Data.List.Relation.Ternary.Appending.Setoid.Properties {c l} (S : Setoid c l) where
11+
module Data.List.Relation.Ternary.Appending.Setoid.Properties
12+
{c l} (S : Setoid c l)
13+
where
1214

1315
open import Data.List.Base as List using (List; [])
1416
import Data.List.Properties as Listₚ
1517
open import Data.List.Relation.Binary.Pointwise.Base using (Pointwise; [])
1618
import Data.List.Relation.Ternary.Appending.Properties as Appendingₚ
17-
open import Data.Product.Base using (_,_)
19+
open import Data.Product using (∃-syntax; _×_; _,_)
20+
open import Function.Base using (id)
21+
open import Relation.Binary.Core using (_⇒_)
1822
open import Relation.Binary.PropositionalEquality.Core using (refl)
23+
open import Relation.Binary.Construct.Composition using (_;_)
1924

25+
open Setoid S renaming (Carrier to A)
26+
open import Relation.Binary.Properties.Setoid S using (≈;≈⇒≈; ≈⇒≈;≈)
2027
open import Data.List.Relation.Ternary.Appending.Setoid S
21-
module S = Setoid S; open S renaming (Carrier to A) using (_≈_)
2228

2329
private
2430
variable
25-
as bs cs : List A
31+
as bs cs ds : List A
2632

2733
------------------------------------------------------------------------
2834
-- Re-exporting existing properties
2935

3036
open Appendingₚ public
31-
hiding (respʳ-≋; respˡ-≋)
37+
using (conicalˡ; conicalʳ)
3238

3339
------------------------------------------------------------------------
3440
-- Proving setoid-specific ones
@@ -44,8 +50,23 @@ open Appendingₚ public
4450

4551
respʳ-≋ : {cs′} Appending as bs cs Pointwise _≈_ cs cs′
4652
Appending as bs cs′
47-
respʳ-≋ = Appendingₚ.respʳ-≋ S.trans S.trans
53+
respʳ-≋ = Appendingₚ.respʳ-≋ trans trans
4854

4955
respˡ-≋ : {as′ bs′} Pointwise _≈_ as′ as Pointwise _≈_ bs′ bs
5056
Appending as bs cs Appending as′ bs′ cs
51-
respˡ-≋ = Appendingₚ.respˡ-≋ S.trans S.trans
57+
respˡ-≋ = Appendingₚ.respˡ-≋ trans trans
58+
59+
through→ :
60+
∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs
61+
∃[ ys ] Appending as bs ys × Pointwise _≈_ ys cs
62+
through→ = Appendingₚ.through→ ≈⇒≈;≈ id
63+
64+
through← :
65+
∃[ ys ] Appending as bs ys × Pointwise _≈_ ys cs
66+
∃[ xs ] Pointwise _≈_ as xs × Appending xs bs cs
67+
through← = Appendingₚ.through← ≈;≈⇒≈ id
68+
69+
assoc→ :
70+
∃[ xs ] Appending as bs xs × Appending xs cs ds
71+
∃[ ys ] Appending bs cs ys × Appending as ys ds
72+
assoc→ = Appendingₚ.assoc→ ≈⇒≈;≈ id ≈;≈⇒≈

src/Relation/Binary/Properties/Setoid.agda

Lines changed: 11 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -9,11 +9,13 @@
99
open import Data.Product.Base using (_,_)
1010
open import Function.Base using (_∘_; id; _$_; flip)
1111
open import Relation.Nullary.Negation.Core using (¬_)
12+
open import Relation.Binary.Core using (_⇒_)
1213
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
1314
open import Relation.Binary.Bundles using (Setoid; Preorder; Poset)
1415
open import Relation.Binary.Definitions
1516
using (Symmetric; _Respectsˡ_; _Respectsʳ_; _Respects₂_)
1617
open import Relation.Binary.Structures using (IsPreorder; IsPartialOrder)
18+
open import Relation.Binary.Construct.Composition using (_;_)
1719

1820
module Relation.Binary.Properties.Setoid {a ℓ} (S : Setoid a ℓ) where
1921

@@ -77,6 +79,15 @@ preorder = record
7779
≉-resp₂ : _≉_ Respects₂ _≈_
7880
≉-resp₂ = ≉-respʳ , ≉-respˡ
7981

82+
------------------------------------------------------------------------
83+
-- Equality is closed under composition
84+
85+
≈;≈⇒≈ : _≈_ ; _≈_ ⇒ _≈_
86+
≈;≈⇒≈ (_ , p , q) = trans p q
87+
88+
≈⇒≈;≈ : _≈_ ⇒ _≈_ ; _≈_
89+
≈⇒≈;≈ q = _ , q , refl
90+
8091
------------------------------------------------------------------------
8192
-- Other properties
8293

0 commit comments

Comments
 (0)