diff --git a/CHANGELOG.md b/CHANGELOG.md index 202b5f4a41..0485996a5c 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -2684,7 +2684,6 @@ Other minor changes toList-++ : toList (xs ++ ys) ≡ toList xs List.++ toList ys - toList-cast : toList (cast eq xs) ≡ toList xs cast-is-id : cast eq xs ≡ xs subst-is-cast : subst (Vec A) eq xs ≡ cast eq xs cast-trans : cast eq₂ (cast eq₁ xs) ≡ cast (trans eq₁ eq₂) xs @@ -2692,8 +2691,20 @@ Other minor changes lookup-cast : lookup (cast eq xs) (Fin.cast eq i) ≡ lookup xs i lookup-cast₁ : lookup (cast eq xs) i ≡ lookup xs (Fin.cast (sym eq) i) lookup-cast₂ : lookup xs (Fin.cast eq i) ≡ lookup (cast (sym eq) xs) i + cast-reverse : cast eq ∘ reverse ≗ reverse ∘ cast eq zipwith-++ : zipWith f (xs ++ ys) (xs' ++ ys') ≡ zipWith f xs xs' ++ zipWith f ys ys' + + ++-assoc : cast eq ((xs ++ ys) ++ zs) ≡ xs ++ (ys ++ zs) + ++-identityʳ : cast eq (xs ++ []) ≡ xs + init-reverse : init ∘ reverse ≗ reverse ∘ tail + last-reverse : last ∘ reverse ≗ head + reverse-++ : cast eq (reverse (xs ++ ys)) ≡ reverse ys ++ reverse xs + + toList-cast : toList (cast eq xs) ≡ toList xs + cast-fromList : cast _ (fromList xs) ≡ fromList ys + fromList-map : cast _ (fromList (List.map f xs)) ≡ map f (fromList xs) + fromList-++ : cast _ (fromList (xs List.++ ys)) ≡ fromList xs ++ fromList ys ``` * Added new proofs in `Data.Vec.Membership.Propositional.Properties`: diff --git a/src/Data/Fin/Properties.agda b/src/Data/Fin/Properties.agda index e8a1cef5dc..0f520afda5 100644 --- a/src/Data/Fin/Properties.agda +++ b/src/Data/Fin/Properties.agda @@ -267,7 +267,7 @@ cast-is-id eq (suc k) = cong suc (cast-is-id (ℕₚ.suc-injective eq) k) subst-is-cast : (eq : m ≡ n) (k : Fin m) → subst Fin eq k ≡ cast eq k subst-is-cast refl k = sym (cast-is-id refl k) -cast-trans : .(eq₁ : m ≡ n) (eq₂ : n ≡ o) (k : Fin m) → +cast-trans : .(eq₁ : m ≡ n) .(eq₂ : n ≡ o) (k : Fin m) → cast eq₂ (cast eq₁ k) ≡ cast (trans eq₁ eq₂) k cast-trans {m = suc _} {n = suc _} {o = suc _} eq₁ eq₂ zero = refl cast-trans {m = suc _} {n = suc _} {o = suc _} eq₁ eq₂ (suc k) = diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda index b6c1f165be..47ff393d0f 100644 --- a/src/Data/Vec/Properties.agda +++ b/src/Data/Vec/Properties.agda @@ -12,9 +12,10 @@ open import Algebra.Definitions open import Data.Bool.Base using (true; false) open import Data.Fin.Base as Fin using (Fin; zero; suc; toℕ; fromℕ<; _↑ˡ_; _↑ʳ_) open import Data.List.Base as List using (List) +import Data.List.Properties as Listₚ open import Data.Nat.Base open import Data.Nat.Properties - using (+-assoc; m≤n⇒m≤1+n; ≤-refl; ≤-trans; suc-injective) + using (+-assoc; m≤n⇒m≤1+n; ≤-refl; ≤-trans; suc-injective; +-comm; +-suc) open import Data.Product.Base as Prod using (_×_; _,_; proj₁; proj₂; <_,_>; uncurry) 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 ------------------------------------------------------------------------ -- cast -toList-cast : ∀ .(eq : m ≡ n) (xs : Vec A m) → toList (cast eq xs) ≡ toList xs -toList-cast {n = zero} eq [] = refl -toList-cast {n = suc _} eq (x ∷ xs) = - cong (x List.∷_) (toList-cast (cong pred eq) xs) - cast-is-id : .(eq : m ≡ m) (xs : Vec A m) → cast eq xs ≡ xs cast-is-id eq [] = refl 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) ++-injective ws xs eq = (++-injectiveˡ ws xs eq , ++-injectiveʳ ws xs eq) +++-assoc : ∀ .(eq : (m + n) + o ≡ m + (n + o)) (xs : Vec A m) (ys : Vec A n) (zs : Vec A o) → + cast eq ((xs ++ ys) ++ zs) ≡ xs ++ (ys ++ zs) +++-assoc eq [] ys zs = cast-is-id eq (ys ++ zs) +++-assoc eq (x ∷ xs) ys zs = cong (x ∷_) (++-assoc (cong pred eq) xs ys zs) + +++-identityʳ : ∀ .(eq : n + zero ≡ n) (xs : Vec A n) → cast eq (xs ++ []) ≡ xs +++-identityʳ eq [] = refl +++-identityʳ eq (x ∷ xs) = cong (x ∷_) (++-identityʳ (cong pred eq) xs) + lookup-++-< : ∀ (xs : Vec A m) (ys : Vec A n) → ∀ i (i