Skip to content

Commit a6da449

Browse files
authored
Add left- and right- Pointwise congruence for _++_ on List (#2426)
* fixes #1131 * response to review comments
1 parent c00030a commit a6da449

File tree

3 files changed

+36
-8
lines changed

3 files changed

+36
-8
lines changed

CHANGELOG.md

+12
Original file line numberDiff line numberDiff line change
@@ -32,6 +32,18 @@ New modules
3232
Additions to existing modules
3333
-----------------------------
3434

35+
* In `Data.List.Relation.Binary.Equality.Setoid`:
36+
```agda
37+
++⁺ʳ : ∀ xs → ys ≋ zs → xs ++ ys ≋ xs ++ zs
38+
++⁺ˡ : ∀ zs → ws ≋ xs → ws ++ zs ≋ xs ++ zs
39+
```
40+
41+
* In `Data.List.Relation.Binary.Pointwise`:
42+
```agda
43+
++⁺ʳ : Reflexive R → ∀ xs → (xs ++_) Preserves (Pointwise R) ⟶ (Pointwise R)
44+
++⁺ˡ : Reflexive R → ∀ zs → (_++ zs) Preserves (Pointwise R) ⟶ (Pointwise R)
45+
```
46+
3547
* New lemma in `Data.Vec.Properties`:
3648
```agda
3749
map-concat : map f (concat xss) ≡ concat (map (map f) xss)

src/Data/List/Relation/Binary/Equality/Setoid.agda

+13-5
Original file line numberDiff line numberDiff line change
@@ -29,6 +29,8 @@ open Setoid S renaming (Carrier to A)
2929
private
3030
variable
3131
p q : Level
32+
ws xs xs′ ys ys′ zs : List A
33+
xss yss : List (List A)
3234

3335
------------------------------------------------------------------------
3436
-- Definition of equality
@@ -113,9 +115,15 @@ foldr⁺ ∙⇔◦ e≈f xs≋ys = PW.foldr⁺ ∙⇔◦ e≈f xs≋ys
113115
------------------------------------------------------------------------
114116
-- _++_
115117

116-
++⁺ : {ws xs ys zs} ws ≋ xs ys ≋ zs ws ++ ys ≋ xs ++ zs
118+
++⁺ : ws ≋ xs ys ≋ zs ws ++ ys ≋ xs ++ zs
117119
++⁺ = PW.++⁺
118120

121+
++⁺ʳ : xs ys ≋ zs xs ++ ys ≋ xs ++ zs
122+
++⁺ʳ xs = PW.++⁺ʳ refl xs
123+
124+
++⁺ˡ : zs ws ≋ xs ws ++ zs ≋ xs ++ zs
125+
++⁺ˡ zs = PW.++⁺ˡ refl zs
126+
119127
++-cancelˡ : xs {ys zs} xs ++ ys ≋ xs ++ zs ys ≋ zs
120128
++-cancelˡ xs = PW.++-cancelˡ xs
121129

@@ -125,7 +133,7 @@ foldr⁺ ∙⇔◦ e≈f xs≋ys = PW.foldr⁺ ∙⇔◦ e≈f xs≋ys
125133
------------------------------------------------------------------------
126134
-- concat
127135

128-
concat⁺ : {xss yss} Pointwise _≋_ xss yss concat xss ≋ concat yss
136+
concat⁺ : Pointwise _≋_ xss yss concat xss ≋ concat yss
129137
concat⁺ = PW.concat⁺
130138

131139
------------------------------------------------------------------------
@@ -146,14 +154,14 @@ module _ {n} {f g : Fin n → A}
146154
module _ {P : Pred A p} (P? : U.Decidable P) (resp : P Respects _≈_)
147155
where
148156

149-
filter⁺ : {xs ys} xs ≋ ys filter P? xs ≋ filter P? ys
157+
filter⁺ : xs ≋ ys filter P? xs ≋ filter P? ys
150158
filter⁺ xs≋ys = PW.filter⁺ P? P? resp (resp ∘ sym) xs≋ys
151159

152160
------------------------------------------------------------------------
153161
-- reverse
154162

155-
ʳ++⁺ : {xs xs′ ys ys′} xs ≋ xs′ ys ≋ ys′ xs ʳ++ ys ≋ xs′ ʳ++ ys′
163+
ʳ++⁺ : xs ≋ xs′ ys ≋ ys′ xs ʳ++ ys ≋ xs′ ʳ++ ys′
156164
ʳ++⁺ = PW.ʳ++⁺
157165

158-
reverse⁺ : {xs ys} xs ≋ ys reverse xs ≋ reverse ys
166+
reverse⁺ : xs ≋ ys reverse xs ≋ reverse ys
159167
reverse⁺ = PW.reverse⁺

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

+11-3
Original file line numberDiff line numberDiff line change
@@ -26,7 +26,7 @@ open import Relation.Nullary hiding (Irrelevant)
2626
import Relation.Nullary.Decidable as Dec using (map′)
2727
open import Relation.Unary as U using (Pred)
2828
open import Relation.Binary.Core renaming (Rel to Rel₂)
29-
open import Relation.Binary.Definitions using (_Respects_; _Respects₂_)
29+
open import Relation.Binary.Definitions using (Reflexive; _Respects_; _Respects₂_)
3030
open import Relation.Binary.Bundles using (Setoid; DecSetoid; Preorder; Poset)
3131
open import Relation.Binary.Structures using (IsEquivalence; IsDecEquivalence; IsPartialOrder; IsPreorder)
3232
open import Relation.Binary.PropositionalEquality.Core as ≡ using (_≡_)
@@ -166,6 +166,15 @@ tabulate⁻ {n = suc n} (x∼y ∷ xs∼ys) (fsuc i) = tabulate⁻ xs∼ys i
166166
++-cancelʳ {xs = xs} (y ∷ ys) [] eq =
167167
contradiction (≡.trans (≡.sym (length-++ (y ∷ ys))) (Pointwise-length eq)) (m≢1+n+m (length xs) ∘ ≡.sym)
168168

169+
module _ (rfl : Reflexive R) where
170+
171+
++⁺ʳ : xs (xs ++_) Preserves (Pointwise R) ⟶ (Pointwise R)
172+
++⁺ʳ xs = ++⁺ (refl rfl)
173+
174+
++⁺ˡ : zs (_++ zs) Preserves (Pointwise R) ⟶ (Pointwise R)
175+
++⁺ˡ zs rs = ++⁺ rs (refl rfl)
176+
177+
169178
------------------------------------------------------------------------
170179
-- concat
171180

@@ -261,8 +270,7 @@ lookup⁺ (_ ∷ Rxys) (fsuc i) = lookup⁺ Rxys i
261270

262271
Pointwise-≡⇒≡ : Pointwise {A = A} _≡_ ⇒ _≡_
263272
Pointwise-≡⇒≡ [] = ≡.refl
264-
Pointwise-≡⇒≡ (≡.refl ∷ xs∼ys) with Pointwise-≡⇒≡ xs∼ys
265-
... | ≡.refl = ≡.refl
273+
Pointwise-≡⇒≡ (≡.refl ∷ xs∼ys) = ≡.cong (_ ∷_) (Pointwise-≡⇒≡ xs∼ys)
266274

267275
≡⇒Pointwise-≡ : _≡_ ⇒ Pointwise {A = A} _≡_
268276
≡⇒Pointwise-≡ ≡.refl = refl ≡.refl

0 commit comments

Comments
 (0)