@@ -12,9 +12,10 @@ open import Algebra.Definitions
12
12
open import Data.Bool.Base using (true; false)
13
13
open import Data.Fin.Base as Fin using (Fin; zero; suc; toℕ; fromℕ<; _↑ˡ_; _↑ʳ_)
14
14
open import Data.List.Base as List using (List)
15
+ import Data.List.Properties as Listₚ
15
16
open import Data.Nat.Base
16
17
open import Data.Nat.Properties
17
- using (+-assoc; m≤n⇒m≤1+n; ≤-refl; ≤-trans; suc-injective)
18
+ using (+-assoc; m≤n⇒m≤1+n; ≤-refl; ≤-trans; suc-injective; +-comm; +-suc )
18
19
open import Data.Product.Base as Prod
19
20
using (_×_; _,_; proj₁; proj₂; <_,_>; uncurry)
20
21
open import Data.Sum.Base using ([_,_]′)
@@ -387,11 +388,6 @@ lookup∘update′ {i = i} {j} i≢j xs y = lookup∘updateAt′ i j i≢j xs
387
388
------------------------------------------------------------------------
388
389
-- cast
389
390
390
- toList-cast : ∀ .(eq : m ≡ n) (xs : Vec A m) → toList (cast eq xs) ≡ toList xs
391
- toList-cast {n = zero} eq [] = refl
392
- toList-cast {n = suc _} eq (x ∷ xs) =
393
- cong (x List.∷_) (toList-cast (cong pred eq) xs)
394
-
395
391
cast-is-id : .(eq : m ≡ m) (xs : Vec A m) → cast eq xs ≡ xs
396
392
cast-is-id eq [] = refl
397
393
cast-is-id eq (x ∷ xs) = cong (x ∷_) (cast-is-id (suc-injective eq) xs)
@@ -488,6 +484,15 @@ toList-map f (x ∷ xs) = cong (f x List.∷_) (toList-map f xs)
488
484
++-injective ws xs eq =
489
485
(++-injectiveˡ ws xs eq , ++-injectiveʳ ws xs eq)
490
486
487
+ ++-assoc : ∀ .(eq : (m + n) + o ≡ m + (n + o)) (xs : Vec A m) (ys : Vec A n) (zs : Vec A o) →
488
+ cast eq ((xs ++ ys) ++ zs) ≡ xs ++ (ys ++ zs)
489
+ ++-assoc eq [] ys zs = cast-is-id eq (ys ++ zs)
490
+ ++-assoc eq (x ∷ xs) ys zs = cong (x ∷_) (++-assoc (cong pred eq) xs ys zs)
491
+
492
+ ++-identityʳ : ∀ .(eq : n + zero ≡ n) (xs : Vec A n) → cast eq (xs ++ []) ≡ xs
493
+ ++-identityʳ eq [] = refl
494
+ ++-identityʳ eq (x ∷ xs) = cong (x ∷_) (++-identityʳ (cong pred eq) xs)
495
+
491
496
lookup-++-< : ∀ (xs : Vec A m) (ys : Vec A n) →
492
497
∀ i (i<m : toℕ i < m) →
493
498
lookup (xs ++ ys) i ≡ lookup xs (Fin.fromℕ< i<m)
@@ -970,6 +975,20 @@ reverse-injective {xs = xs} {ys} eq = begin
970
975
reverse (reverse ys) ≡⟨ reverse-involutive ys ⟩
971
976
ys ∎
972
977
978
+ -- init and last of reverse
979
+
980
+ init-reverse : init ∘ reverse ≗ reverse ∘ tail {A = A} {n = n}
981
+ init-reverse (x ∷ xs) = begin
982
+ init (reverse (x ∷ xs)) ≡⟨ cong init (reverse-∷ x xs) ⟩
983
+ init (reverse xs ∷ʳ x) ≡⟨ init-∷ʳ x (reverse xs) ⟩
984
+ reverse xs ∎
985
+
986
+ last-reverse : last ∘ reverse ≗ head {A = A} {n = n}
987
+ last-reverse (x ∷ xs) = begin
988
+ last (reverse (x ∷ xs)) ≡⟨ cong last (reverse-∷ x xs) ⟩
989
+ last (reverse xs ∷ʳ x) ≡⟨ last-∷ʳ x (reverse xs) ⟩
990
+ x ∎
991
+
973
992
-- map and reverse
974
993
975
994
map-reverse : ∀ (f : A → B) (xs : Vec A n) →
@@ -983,6 +1002,39 @@ map-reverse f (x ∷ xs) = begin
983
1002
reverse (f x ∷ map f xs) ≡⟨⟩
984
1003
reverse (map f (x ∷ xs)) ∎
985
1004
1005
+ -- append and reverse
1006
+
1007
+ reverse-++ : ∀ .(eq : m + n ≡ n + m) (xs : Vec A m) (ys : Vec A n) →
1008
+ 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₂
1027
+
1028
+ cast-reverse : ∀ .(eq : m ≡ n) → cast eq ∘ reverse {A = A} {n = m} ≗ reverse ∘ cast eq
1029
+ 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)) ∎
1037
+
986
1038
------------------------------------------------------------------------
987
1039
-- _ʳ++_
988
1040
@@ -1161,6 +1213,29 @@ toList∘fromList : (xs : List A) → toList (fromList xs) ≡ xs
1161
1213
toList∘fromList List.[] = refl
1162
1214
toList∘fromList (x List.∷ xs) = cong (x List.∷_) (toList∘fromList xs)
1163
1215
1216
+ toList-cast : ∀ .(eq : m ≡ n) (xs : Vec A m) → toList (cast eq xs) ≡ toList xs
1217
+ toList-cast {n = zero} eq [] = refl
1218
+ toList-cast {n = suc _} eq (x ∷ xs) =
1219
+ cong (x List.∷_) (toList-cast (cong pred eq) xs)
1220
+
1221
+ cast-fromList : ∀ {xs ys : List A} (eq : xs ≡ ys) →
1222
+ cast (cong List.length eq) (fromList xs) ≡ fromList ys
1223
+ cast-fromList {xs = List.[]} {ys = List.[]} eq = refl
1224
+ cast-fromList {xs = x List.∷ xs} {ys = y List.∷ ys} eq =
1225
+ let x≡y , xs≡ys = Listₚ.∷-injective eq in begin
1226
+ x ∷ cast (cong (pred ∘ List.length) eq) (fromList xs) ≡⟨ cong (_ ∷_) (cast-fromList xs≡ys) ⟩
1227
+ x ∷ fromList ys ≡⟨ cong (_∷ _) x≡y ⟩
1228
+ y ∷ fromList ys ∎
1229
+
1230
+ fromList-map : ∀ (f : A → B) (xs : List A) →
1231
+ cast (Listₚ.length-map f xs) (fromList (List.map f xs)) ≡ map f (fromList xs)
1232
+ fromList-map f List.[] = refl
1233
+ fromList-map f (x List.∷ xs) = cong (f x ∷_) (fromList-map f xs)
1234
+
1235
+ fromList-++ : ∀ (xs : List A) {ys : List A} →
1236
+ cast (Listₚ.length-++ xs) (fromList (xs List.++ ys)) ≡ fromList xs ++ fromList ys
1237
+ fromList-++ List.[] {ys} = cast-is-id refl (fromList ys)
1238
+ fromList-++ (x List.∷ xs) {ys} = cong (x ∷_) (fromList-++ xs)
1164
1239
1165
1240
------------------------------------------------------------------------
1166
1241
-- DEPRECATED NAMES
0 commit comments