Skip to content

Commit db26bb9

Browse files
committed
Add combinators for equational reasoning of vector
1 parent fd541d0 commit db26bb9

File tree

3 files changed

+227
-25
lines changed

3 files changed

+227
-25
lines changed

CHANGELOG.md

Lines changed: 13 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -1765,6 +1765,11 @@ New modules
17651765
Function.Indexed.Bundles
17661766
```
17671767

1768+
* Combinators for propositional equational reasoning on vectors with different indices
1769+
```
1770+
Data.Vec.Relation.Binary.Reasoning.Propositional
1771+
```
1772+
17681773
Other minor changes
17691774
-------------------
17701775

@@ -2673,6 +2678,7 @@ Other minor changes
26732678
last-∷ʳ : last (xs ∷ʳ x) ≡ x
26742679
cast-∷ʳ : cast eq (xs ∷ʳ x) ≡ (cast (cong pred eq) xs) ∷ʳ x
26752680
++-∷ʳ : cast eq ((xs ++ ys) ∷ʳ z) ≡ xs ++ (ys ∷ʳ z)
2681+
∷ʳ-++ : cast eq ((xs ∷ʳ a) ++ ys) ≡ xs ++ (a ∷ ys)
26762682
26772683
reverse-∷ : reverse (x ∷ xs) ≡ reverse xs ∷ʳ x
26782684
reverse-involutive : Involutive _≡_ reverse
@@ -2692,6 +2698,8 @@ Other minor changes
26922698
lookup-cast₁ : lookup (cast eq xs) i ≡ lookup xs (Fin.cast (sym eq) i)
26932699
lookup-cast₂ : lookup xs (Fin.cast eq i) ≡ lookup (cast (sym eq) xs) i
26942700
cast-reverse : cast eq ∘ reverse ≗ reverse ∘ cast eq
2701+
cast-++ˡ : cast (cong (_+ n) eq) (xs ++ ys) ≡ cast eq xs ++ ys
2702+
cast-++ʳ : cast (cong (m +_) eq) (xs ++ ys) ≡ xs ++ cast eq ys
26952703
26962704
zipwith-++ : zipWith f (xs ++ ys) (xs' ++ ys') ≡ zipWith f xs xs' ++ zipWith f ys ys'
26972705
@@ -2705,6 +2713,11 @@ Other minor changes
27052713
cast-fromList : cast _ (fromList xs) ≡ fromList ys
27062714
fromList-map : cast _ (fromList (List.map f xs)) ≡ map f (fromList xs)
27072715
fromList-++ : cast _ (fromList (xs List.++ ys)) ≡ fromList xs ++ fromList ys
2716+
fromList-reverse : cast (Listₚ.length-reverse xs) (fromList (List.reverse xs)) ≡ reverse (fromList xs)
2717+
2718+
∷-ʳ++ : cast eq ((a ∷ xs) ʳ++ ys) ≡ xs ʳ++ (a ∷ ys)
2719+
++-ʳ++ : cast eq ((xs ++ ys) ʳ++ zs) ≡ ys ʳ++ (xs ʳ++ zs)
2720+
ʳ++-ʳ++ : cast eq ((xs ʳ++ ys) ʳ++ zs) ≡ ys ʳ++ (xs ++ zs)
27082721
```
27092722

27102723
* Added new proofs in `Data.Vec.Membership.Propositional.Properties`:

src/Data/Vec/Properties.agda

Lines changed: 75 additions & 25 deletions
Original file line numberDiff line numberDiff line change
@@ -21,6 +21,7 @@ open import Data.Product.Base as Prod
2121
open import Data.Sum.Base using ([_,_]′)
2222
open import Data.Sum.Properties using ([,]-map)
2323
open import Data.Vec.Base
24+
open import Data.Vec.Relation.Binary.Reasoning.Propositional as VecReasoning renaming (begin_ to begin′_; _∎ to _∎′)
2425
open import Function.Base
2526
-- open import Function.Inverse using (_↔_; inverse)
2627
open import Function.Bundles using (_↔_; mk↔′)
@@ -493,6 +494,16 @@ toList-map f (x ∷ xs) = cong (f x List.∷_) (toList-map f xs)
493494
++-identityʳ eq [] = refl
494495
++-identityʳ eq (x ∷ xs) = cong (x ∷_) (++-identityʳ (cong pred eq) xs)
495496

497+
cast-++ˡ : .(eq : m ≡ o) (xs : Vec A m) {ys : Vec A n}
498+
cast (cong (_+ n) eq) (xs ++ ys) ≡ cast eq xs ++ ys
499+
cast-++ˡ {o = zero} eq [] {ys} = cast-is-id refl (cast eq [] ++ ys)
500+
cast-++ˡ {o = suc o} eq (x ∷ xs) {ys} = cong (x ∷_) (cast-++ˡ (cong pred eq) xs)
501+
502+
cast-++ʳ : .(eq : n ≡ o) (xs : Vec A m) {ys : Vec A n}
503+
cast (cong (m +_) eq) (xs ++ ys) ≡ xs ++ cast eq ys
504+
cast-++ʳ {m = zero} eq [] {ys} = refl
505+
cast-++ʳ {m = suc m} eq (x ∷ xs) {ys} = cong (x ∷_) (cast-++ʳ eq xs)
506+
496507
lookup-++-< : (xs : Vec A m) (ys : Vec A n)
497508
i (i<m : toℕ i < m)
498509
lookup (xs ++ ys) i ≡ lookup xs (Fin.fromℕ< i<m)
@@ -927,6 +938,11 @@ cast-∷ʳ {m = suc m} eq x (y ∷ xs) = cong (y ∷_) (cast-∷ʳ (cong pred eq
927938
++-∷ʳ {m = zero} eq z [] (y ∷ ys) = cong (y ∷_) (++-∷ʳ refl z [] ys)
928939
++-∷ʳ {m = suc m} eq z (x ∷ xs) ys = cong (x ∷_) (++-∷ʳ (cong pred eq) z xs ys)
929940

941+
∷ʳ-++ : .(eq : (suc n) + m ≡ n + suc m) a (xs : Vec A n) {ys}
942+
cast eq ((xs ∷ʳ a) ++ ys) ≡ xs ++ (a ∷ ys)
943+
∷ʳ-++ eq a [] {ys} = cong (a ∷_) (cast-is-id (cong pred eq) ys)
944+
∷ʳ-++ eq a (x ∷ xs) {ys} = cong (x ∷_) (∷ʳ-++ (cong pred eq) a xs)
945+
930946
------------------------------------------------------------------------
931947
-- reverse
932948

@@ -1006,34 +1022,24 @@ map-reverse f (x ∷ xs) = begin
10061022

10071023
reverse-++ : .(eq : m + n ≡ n + m) (xs : Vec A m) (ys : Vec A n)
10081024
cast eq (reverse (xs ++ ys)) ≡ reverse ys ++ reverse xs
1009-
reverse-++ {m = zero} {n = n} eq [] ys = begin
1010-
cast _ (reverse ys) ≡˘⟨ cong (cast eq) (++-identityʳ (sym eq) (reverse ys)) ⟩
1011-
cast _ (cast _ (reverse ys ++ [])) ≡⟨ cast-trans (sym eq) eq (reverse ys ++ []) ⟩
1012-
cast _ (reverse ys ++ []) ≡⟨ cast-is-id (trans (sym eq) eq) (reverse ys ++ []) ⟩
1013-
reverse ys ++ [] ≡⟨⟩
1014-
reverse ys ++ reverse [] ∎
1015-
reverse-++ {m = suc m} {n = n} eq (x ∷ xs) ys = begin
1016-
cast eq (reverse (x ∷ xs ++ ys)) ≡⟨ cong (cast eq) (reverse-∷ x (xs ++ ys)) ⟩
1017-
cast eq (reverse (xs ++ ys) ∷ʳ x) ≡˘⟨ cast-trans eq₂ eq₁ (reverse (xs ++ ys) ∷ʳ x) ⟩
1018-
(cast eq₁ ∘ cast eq₂) (reverse (xs ++ ys) ∷ʳ x) ≡⟨ cong (cast eq₁) (cast-∷ʳ _ x (reverse (xs ++ ys))) ⟩
1019-
cast eq₁ ((cast eq₃ (reverse (xs ++ ys))) ∷ʳ x) ≡⟨ cong (cast eq₁) (cong (_∷ʳ x) (reverse-++ _ xs ys)) ⟩
1020-
cast eq₁ ((reverse ys ++ reverse xs) ∷ʳ x) ≡⟨ ++-∷ʳ _ x (reverse ys) (reverse xs) ⟩
1021-
reverse ys ++ (reverse xs ∷ʳ x) ≡˘⟨ cong (reverse ys ++_) (reverse-∷ x xs) ⟩
1022-
reverse ys ++ (reverse (x ∷ xs)) ∎
1023-
where
1024-
eq₁ = sym (+-suc n m)
1025-
eq₂ = cong suc (+-comm m n)
1026-
eq₃ = cong pred eq₂
1025+
reverse-++ {m = zero} {n = n} eq [] ys = ≈-sym (++-identityʳ (sym eq) (reverse ys))
1026+
reverse-++ {m = suc m} {n = n} eq (x ∷ xs) ys = begin′
1027+
reverse (x ∷ xs ++ ys) ≂⟨ reverse-∷ x (xs ++ ys) ⟩
1028+
reverse (xs ++ ys) ∷ʳ x ≈⟨ cast-∷ʳ (cong suc (+-comm m n)) x (reverse (xs ++ ys))
1029+
≈cong[ (_∷ʳ x) ] reverse-++ _ xs ys ⟩
1030+
(reverse ys ++ reverse xs) ∷ʳ x ≈⟨ ++-∷ʳ (sym (+-suc n m)) x (reverse ys) (reverse xs) ⟩
1031+
reverse ys ++ (reverse xs ∷ʳ x) ≂˘⟨ cong (reverse ys ++_) (reverse-∷ x xs) ⟩
1032+
reverse ys ++ (reverse (x ∷ xs)) ∎′
10271033

10281034
cast-reverse : .(eq : m ≡ n) cast eq ∘ reverse {A = A} {n = m} ≗ reverse ∘ cast eq
10291035
cast-reverse {n = zero} eq [] = refl
1030-
cast-reverse {n = suc n} eq (x ∷ xs) = begin
1031-
cast eq (reverse (x ∷ xs)) ≡⟨ cong (cast eq) (reverse-∷ x xs)
1032-
cast eq (reverse xs ∷ʳ x) ⟨ cast-∷ʳ eq x (reverse xs)
1033-
(cast (cong pred eq) (reverse xs)) ∷ʳ x ≡⟨ cong (_∷ʳ x) (cast-reverse (cong pred eq) xs)
1034-
(reverse (cast (cong pred eq) xs)) ∷ʳ x ˘⟨ reverse-∷ x (cast (cong pred eq) xs) ⟩
1035-
reverse (x ∷ cast (cong pred eq) xs) ⟨⟩
1036-
reverse (cast eq (x ∷ xs))
1036+
cast-reverse {n = suc n} eq (x ∷ xs) = begin
1037+
reverse (x ∷ xs) ≂⟨ reverse-∷ x xs ⟩
1038+
reverse xs ∷ʳ x ⟨ cast-∷ʳ eq x (reverse xs)
1039+
cong[ (_∷ʳ x) ] cast-reverse (cong pred eq) xs ⟩
1040+
reverse (cast _ xs) ∷ʳ x ˘⟨ reverse-∷ x (cast (cong pred eq) xs) ⟩
1041+
reverse (x ∷ cast _ xs) ⟨⟩
1042+
reverse (cast eq (x ∷ xs)) ∎′
10371043

10381044
------------------------------------------------------------------------
10391045
-- _ʳ++_
@@ -1062,6 +1068,38 @@ map-ʳ++ {ys = ys} f xs = begin
10621068
reverse (map f xs) ++ map f ys ≡˘⟨ unfold-ʳ++ (map f xs) (map f ys) ⟩
10631069
map f xs ʳ++ map f ys ∎
10641070

1071+
∷-ʳ++ : .(eq : (suc m) + n ≡ m + suc n) a (xs : Vec A m) {ys}
1072+
cast eq ((a ∷ xs) ʳ++ ys) ≡ xs ʳ++ (a ∷ ys)
1073+
∷-ʳ++ eq a xs {ys} = begin′
1074+
(a ∷ xs) ʳ++ ys ≂⟨ unfold-ʳ++ (a ∷ xs) ys ⟩
1075+
reverse (a ∷ xs) ++ ys ≂⟨ cong (_++ ys) (reverse-∷ a xs) ⟩
1076+
(reverse xs ∷ʳ a) ++ ys ≈⟨ ∷ʳ-++ eq a (reverse xs) ⟩
1077+
reverse xs ++ (a ∷ ys) ≂˘⟨ unfold-ʳ++ xs (a ∷ ys) ⟩
1078+
xs ʳ++ (a ∷ ys) ∎′
1079+
1080+
++-ʳ++ : .(eq : m + n + o ≡ n + (m + o)) (xs : Vec A m) {ys : Vec A n} {zs : Vec A o}
1081+
cast eq ((xs ++ ys) ʳ++ zs) ≡ ys ʳ++ (xs ʳ++ zs)
1082+
++-ʳ++ {m = m} {n} {o} eq xs {ys} {zs} = begin′
1083+
((xs ++ ys) ʳ++ zs) ≂⟨ unfold-ʳ++ (xs ++ ys) zs ⟩
1084+
reverse (xs ++ ys) ++ zs ≈⟨ cast-++ˡ (+-comm m n) (reverse (xs ++ ys))
1085+
≈cong[ (_++ zs) ] reverse-++ (+-comm m n) xs ys ⟩
1086+
(reverse ys ++ reverse xs) ++ zs ≈⟨ ++-assoc (trans (cong (_+ o) (+-comm n m)) eq) (reverse ys) (reverse xs) zs ⟩
1087+
reverse ys ++ (reverse xs ++ zs) ≂˘⟨ cong (reverse ys ++_) (unfold-ʳ++ xs zs) ⟩
1088+
reverse ys ++ (xs ʳ++ zs) ≂˘⟨ unfold-ʳ++ ys (xs ʳ++ zs) ⟩
1089+
ys ʳ++ (xs ʳ++ zs) ∎′
1090+
1091+
ʳ++-ʳ++ : .(eq : (m + n) + o ≡ n + (m + o)) (xs : Vec A m) {ys : Vec A n} {zs}
1092+
cast eq ((xs ʳ++ ys) ʳ++ zs) ≡ ys ʳ++ (xs ++ zs)
1093+
ʳ++-ʳ++ {m = m} {n} {o} eq xs {ys} {zs} = begin′
1094+
(xs ʳ++ ys) ʳ++ zs ≂⟨ cong (_ʳ++ zs) (unfold-ʳ++ xs ys) ⟩
1095+
(reverse xs ++ ys) ʳ++ zs ≂⟨ unfold-ʳ++ (reverse xs ++ ys) zs ⟩
1096+
reverse (reverse xs ++ ys) ++ zs ≈⟨ cast-++ˡ (+-comm m n) (reverse (reverse xs ++ ys))
1097+
≈cong[ (_++ zs) ] reverse-++ (+-comm m n) (reverse xs) ys ⟩
1098+
(reverse ys ++ reverse (reverse xs)) ++ zs ≂⟨ cong ((_++ zs) ∘ (reverse ys ++_)) (reverse-involutive xs) ⟩
1099+
(reverse ys ++ xs) ++ zs ≈⟨ ++-assoc (+-assoc n m o) (reverse ys) xs zs ⟩
1100+
reverse ys ++ (xs ++ zs) ≂˘⟨ unfold-ʳ++ ys (xs ++ zs) ⟩
1101+
ys ʳ++ (xs ++ zs) ∎′
1102+
10651103
------------------------------------------------------------------------
10661104
-- sum
10671105

@@ -1237,6 +1275,18 @@ fromList-++ : ∀ (xs : List A) {ys : List A} →
12371275
fromList-++ List.[] {ys} = cast-is-id refl (fromList ys)
12381276
fromList-++ (x List.∷ xs) {ys} = cong (x ∷_) (fromList-++ xs)
12391277

1278+
fromList-reverse : (xs : List A) cast (Listₚ.length-reverse xs) (fromList (List.reverse xs)) ≡ reverse (fromList xs)
1279+
fromList-reverse List.[] = refl
1280+
fromList-reverse (x List.∷ xs) = begin′
1281+
fromList (List.reverse (x List.∷ xs)) ≈⟨ cast-fromList (Listₚ.ʳ++-defn xs) ⟩
1282+
fromList (List.reverse xs List.++ List.[ x ]) ≈⟨ fromList-++ (List.reverse xs) ⟩
1283+
fromList (List.reverse xs) ++ [ x ] ≈˘⟨ unfold-∷ʳ (+-comm 1 _) x (fromList (List.reverse xs)) ⟩
1284+
fromList (List.reverse xs) ∷ʳ x ≈⟨ cast-∷ʳ (cong suc (Listₚ.length-reverse xs)) _ _
1285+
≈cong[ (_∷ʳ x) ] fromList-reverse xs ⟩
1286+
reverse (fromList xs) ∷ʳ x ≂˘⟨ reverse-∷ x (fromList xs) ⟩
1287+
reverse (x ∷ fromList xs) ≈⟨⟩
1288+
reverse (fromList (x List.∷ xs)) ∎′
1289+
12401290
------------------------------------------------------------------------
12411291
-- DEPRECATED NAMES
12421292
------------------------------------------------------------------------
Lines changed: 139 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,139 @@
1+
------------------------------------------------------------------------
2+
-- The Agda standard library
3+
--
4+
-- The basic code for equational reasoning with displayed propositional equality over vectors
5+
------------------------------------------------------------------------
6+
7+
{-# OPTIONS --cubical-compatible --safe #-}
8+
9+
module Data.Vec.Relation.Binary.Reasoning.Propositional {a} {A : Set a} where
10+
11+
open import Data.Nat.Base using (ℕ; zero; suc)
12+
open import Data.Nat.Properties using (suc-injective)
13+
open import Data.Vec.Base
14+
open import Relation.Binary.Core using (REL; _⇒_)
15+
open import Relation.Binary.Definitions using (Sym; Trans)
16+
open import Relation.Binary.PropositionalEquality.Core
17+
using (_≡_; refl; trans; sym; cong; module ≡-Reasoning)
18+
19+
private
20+
variable
21+
l m n o :
22+
xs ys zs : Vec A n
23+
24+
-- Duplicate definition of cast-is-id and cast-trans to avoid circular dependency
25+
private
26+
cast-is-id : .(eq : m ≡ m) (xs : Vec A m) cast eq xs ≡ xs
27+
cast-is-id eq [] = refl
28+
cast-is-id eq (x ∷ xs) = cong (x ∷_) (cast-is-id (suc-injective eq) xs)
29+
30+
cast-trans : .(eq₁ : m ≡ n) .(eq₂ : n ≡ o) (xs : Vec A m)
31+
cast eq₂ (cast eq₁ xs) ≡ cast (trans eq₁ eq₂) xs
32+
cast-trans {m = zero} {n = zero} {o = zero} eq₁ eq₂ [] = refl
33+
cast-trans {m = suc _} {n = suc _} {o = suc _} eq₁ eq₂ (x ∷ xs) =
34+
cong (x ∷_) (cast-trans (suc-injective eq₁) (suc-injective eq₂) xs)
35+
36+
37+
38+
≈-by : .(eq : n ≡ m) REL (Vec A n) (Vec A m) _
39+
≈-by eq xs ys = (cast eq xs ≡ ys)
40+
41+
infix 3 ≈-by
42+
syntax ≈-by n≡m xs ys = xs ≈[ n≡m ] ys
43+
44+
----------------------------------------------------------------
45+
-- ≈-by is ‘reflexive’, ‘symmetric’ and ‘transitive’
46+
47+
≈-reflexive : {n} _≡_ ⇒ ≈-by {n} refl
48+
≈-reflexive {x = x} eq = trans (cast-is-id refl x) eq
49+
50+
≈-sym : .{m≡n : m ≡ n} Sym (≈-by m≡n) (≈-by (sym m≡n))
51+
≈-sym {m≡n = m≡n} {xs} {ys} xs≈ys = begin
52+
cast (sym m≡n) ys ≡˘⟨ cong (cast (sym m≡n)) xs≈ys ⟩
53+
cast (sym m≡n) (cast m≡n xs) ≡⟨ cast-trans m≡n (sym m≡n) xs ⟩
54+
cast (trans m≡n (sym m≡n)) xs ≡⟨ cast-is-id (trans m≡n (sym m≡n)) xs ⟩
55+
xs ∎
56+
where open ≡-Reasoning
57+
58+
≈-trans : .{m≡n : m ≡ n} .{n≡o : n ≡ o} Trans (≈-by m≡n) (≈-by n≡o) (≈-by (trans m≡n n≡o))
59+
≈-trans {m≡n = m≡n} {n≡o} {xs} {ys} {zs} xs≈ys ys≈zs = begin
60+
cast (trans m≡n n≡o) xs ≡˘⟨ cast-trans m≡n n≡o xs ⟩
61+
cast n≡o (cast m≡n xs) ≡⟨ cong (cast n≡o) xs≈ys ⟩
62+
cast n≡o ys ≡⟨ ys≈zs ⟩
63+
zs ∎
64+
where open ≡-Reasoning
65+
66+
------------------------------------------------------------------------
67+
-- Reasoning combinators
68+
69+
begin_ : .{m≡n : m ≡ n} {xs : Vec A m} {ys} xs ≈[ m≡n ] ys cast m≡n xs ≡ ys
70+
begin xs≈ys = xs≈ys
71+
72+
_∎ : (xs : Vec A n) cast refl xs ≡ xs
73+
_∎ xs = ≈-reflexive refl
74+
75+
_≈⟨⟩_ : .{m≡n : m ≡ n} (xs : Vec A m) {ys} (xs ≈[ m≡n ] ys) (xs ≈[ m≡n ] ys)
76+
xs ≈⟨⟩ xs≈ys = xs≈ys
77+
78+
-- composition of _≈[_]_
79+
step-≈ : .{m≡n : m ≡ n}.{m≡o : m ≡ o} (xs : Vec A m) {ys : Vec A n} {zs : Vec A o}
80+
(ys ≈[ trans (sym m≡n) m≡o ] zs) (xs ≈[ m≡n ] ys) (xs ≈[ m≡o ] zs)
81+
step-≈ xs ys≈zs xs≈ys = ≈-trans xs≈ys ys≈zs
82+
83+
-- composition of the equality type on the right-hand side of _≈[_]_, or escaping to ordinary _≡_
84+
step-≃ : .{m≡n : m ≡ n} (xs : Vec A m) {ys zs} (ys ≡ zs) (xs ≈[ m≡n ] ys) (xs ≈[ m≡n ] zs)
85+
step-≃ xs ys≡zs xs≈ys = ≈-trans xs≈ys (≈-reflexive ys≡zs)
86+
87+
-- composition of the equality type on the left-hand side of _≈[_]_
88+
step-≂ : .{m≡n : m ≡ n} (xs : Vec A m) {ys zs} (ys ≈[ m≡n ] zs) (xs ≡ ys) (xs ≈[ m≡n ] zs)
89+
step-≂ xs ys≈zs xs≡ys = ≈-trans (≈-reflexive xs≡ys) ys≈zs
90+
91+
-- `cong` after a `_≈[_]_` step that exposes the `cast` to the `cong` operation
92+
≈-cong : .{l≡o : l ≡ o} .{m≡n : m ≡ n} {xs : Vec A m} {ys zs} (f : Vec A o Vec A n)
93+
(ys ≈[ l≡o ] zs) (xs ≈[ m≡n ] f (cast l≡o ys)) (xs ≈[ m≡n ] f zs)
94+
≈-cong f ys≈zs xs≈fys = trans xs≈fys (cong f ys≈zs)
95+
96+
97+
-- symmetric version of each of the operator
98+
step-≈˘ : .{n≡m : n ≡ m}.{m≡o : m ≡ o} (xs : Vec A m) {ys : Vec A n} {zs : Vec A o}
99+
(ys ≈[ trans n≡m m≡o ] zs) (ys ≈[ n≡m ] xs) (xs ≈[ m≡o ] zs)
100+
step-≈˘ xs ys≈zs ys≈xs = ≈-trans (≈-sym ys≈xs) ys≈zs
101+
102+
step-≃˘ : .{m≡n : m ≡ n} (xs : Vec A m) {ys zs} (ys ≡ zs) (ys ≈[ sym m≡n ] xs) (xs ≈[ m≡n ] zs)
103+
step-≃˘ xs ys≡zs ys≈xs = trans (≈-sym ys≈xs) ys≡zs
104+
105+
step-≂˘ : .{m≡n : m ≡ n} (xs : Vec A m) {ys zs} (ys ≈[ m≡n ] zs) (ys ≡ xs) (xs ≈[ m≡n ] zs)
106+
step-≂˘ xs ys≈zs ys≡xs = step-≂ xs ys≈zs (sym ys≡xs)
107+
108+
109+
----------------------------------------------------------------
110+
-- convenient syntax for ‘equational’ reasoning
111+
112+
infix 1 begin_
113+
infixr 2 step-≃ step-≂ step-≃˘ step-≂˘ step-≈ step-≈˘ _≈⟨⟩_ ≈-cong
114+
infix 3 _∎
115+
116+
syntax step-≃ xs ys≡zs xs≈ys = xs ≃⟨ xs≈ys ⟩ ys≡zs
117+
syntax step-≃˘ xs ys≡zs xs≈ys = xs ≃˘⟨ xs≈ys ⟩ ys≡zs
118+
syntax step-≂ xs ys≈zs xs≡ys = xs ≂⟨ xs≡ys ⟩ ys≈zs
119+
syntax step-≂˘ xs ys≈zs ys≡xs = xs ≂˘⟨ ys≡xs ⟩ ys≈zs
120+
syntax step-≈ xs ys≈zs xs≈ys = xs ≈⟨ xs≈ys ⟩ ys≈zs
121+
syntax step-≈˘ xs ys≈zs ys≈xs = xs ≈˘⟨ ys≈xs ⟩ ys≈zs
122+
syntax ≈-cong f ys≈zs xs≈fys = xs≈fys ≈cong[ f ] ys≈zs
123+
124+
{-
125+
-- An equational reasoning example, demonstrating nested uses of the cong operator
126+
127+
cast-++ˡ : ∀ .(eq : n ≡ o) (xs : Vec A n) {ys : Vec A m} →
128+
cast (cong (_+ m) eq) (xs ++ ys) ≡ cast eq xs ++ ys
129+
130+
example : ∀ .(eq : (m + 1) + n ≡ n + suc m) a (xs : Vec A m) (ys : Vec A n) →
131+
cast eq (reverse ((xs ++ [ a ]) ++ ys)) ≡ ys ʳ++ reverse (xs ∷ʳ a)
132+
example {m = m} {n} eq a xs ys = begin
133+
reverse ((xs ++ [ a ]) ++ ys) ≈˘⟨ cast-reverse (cong (_+ n) (ℕₚ.+-comm 1 m)) ((xs ∷ʳ a) ++ ys)
134+
≈cong[ reverse ] cast-++ˡ (ℕₚ.+-comm 1 m) (xs ∷ʳ a)
135+
≈cong[ (_++ ys) ] unfold-∷ʳ _ a xs ⟩
136+
reverse ((xs ∷ʳ a) ++ ys) ≈⟨ reverse-++ (ℕₚ.+-comm (suc m) n) (xs ∷ʳ a) ys ⟩
137+
reverse ys ++ reverse (xs ∷ʳ a) ≂˘⟨ unfold-ʳ++ ys (reverse (xs ∷ʳ a)) ⟩
138+
ys ʳ++ reverse (xs ∷ʳ a) ∎
139+
-}

0 commit comments

Comments
 (0)