@@ -21,6 +21,7 @@ open import Data.Product.Base as Prod
21
21
open import Data.Sum.Base using ([_,_]′)
22
22
open import Data.Sum.Properties using ([,]-map)
23
23
open import Data.Vec.Base
24
+ open import Data.Vec.Relation.Binary.Reasoning.Propositional as VecReasoning renaming (begin_ to begin′_; _∎ to _∎′)
24
25
open import Function.Base
25
26
-- open import Function.Inverse using (_↔_; inverse)
26
27
open import Function.Bundles using (_↔_; mk↔′)
@@ -493,6 +494,16 @@ toList-map f (x ∷ xs) = cong (f x List.∷_) (toList-map f xs)
493
494
++-identityʳ eq [] = refl
494
495
++-identityʳ eq (x ∷ xs) = cong (x ∷_) (++-identityʳ (cong pred eq) xs)
495
496
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
+
496
507
lookup-++-< : ∀ (xs : Vec A m) (ys : Vec A n) →
497
508
∀ i (i<m : toℕ i < m) →
498
509
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
927
938
++-∷ʳ {m = zero} eq z [] (y ∷ ys) = cong (y ∷_) (++-∷ʳ refl z [] ys)
928
939
++-∷ʳ {m = suc m} eq z (x ∷ xs) ys = cong (x ∷_) (++-∷ʳ (cong pred eq) z xs ys)
929
940
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
+
930
946
------------------------------------------------------------------------
931
947
-- reverse
932
948
@@ -1006,34 +1022,24 @@ map-reverse f (x ∷ xs) = begin
1006
1022
1007
1023
reverse-++ : ∀ .(eq : m + n ≡ n + m) (xs : Vec A m) (ys : Vec A n) →
1008
1024
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)) ∎′
1027
1033
1028
1034
cast-reverse : ∀ .(eq : m ≡ n) → cast eq ∘ reverse {A = A} {n = m} ≗ reverse ∘ cast eq
1029
1035
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)) ∎′
1037
1043
1038
1044
------------------------------------------------------------------------
1039
1045
-- _ʳ++_
@@ -1062,6 +1068,38 @@ map-ʳ++ {ys = ys} f xs = begin
1062
1068
reverse (map f xs) ++ map f ys ≡˘⟨ unfold-ʳ++ (map f xs) (map f ys) ⟩
1063
1069
map f xs ʳ++ map f ys ∎
1064
1070
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
+
1065
1103
------------------------------------------------------------------------
1066
1104
-- sum
1067
1105
@@ -1239,6 +1277,18 @@ fromList-++ : ∀ (xs : List A) {ys : List A} →
1239
1277
fromList-++ List.[] {ys} = cast-is-id refl (fromList ys)
1240
1278
fromList-++ (x List.∷ xs) {ys} = cong (x ∷_) (fromList-++ xs)
1241
1279
1280
+ fromList-reverse : (xs : List A) → cast (Listₚ.length-reverse xs) (fromList (List.reverse xs)) ≡ reverse (fromList xs)
1281
+ fromList-reverse List.[] = refl
1282
+ fromList-reverse (x List.∷ xs) = begin′
1283
+ fromList (List.reverse (x List.∷ xs)) ≈⟨ cast-fromList (Listₚ.ʳ++-defn xs) ⟩
1284
+ fromList (List.reverse xs List.++ List.[ x ]) ≈⟨ fromList-++ (List.reverse xs) ⟩
1285
+ fromList (List.reverse xs) ++ [ x ] ≈˘⟨ unfold-∷ʳ (+-comm 1 _) x (fromList (List.reverse xs)) ⟩
1286
+ fromList (List.reverse xs) ∷ʳ x ≈⟨ cast-∷ʳ (cong suc (Listₚ.length-reverse xs)) _ _
1287
+ ≈cong[ (_∷ʳ x) ] fromList-reverse xs ⟩
1288
+ reverse (fromList xs) ∷ʳ x ≂˘⟨ reverse-∷ x (fromList xs) ⟩
1289
+ reverse (x ∷ fromList xs) ≈⟨⟩
1290
+ reverse (fromList (x List.∷ xs)) ∎′
1291
+
1242
1292
------------------------------------------------------------------------
1243
1293
-- DEPRECATED NAMES
1244
1294
------------------------------------------------------------------------
0 commit comments