diff --git a/CHANGELOG.md b/CHANGELOG.md index c636c2f1bb..1983928857 100644 --- a/CHANGELOG.md +++ b/CHANGELOG.md @@ -545,6 +545,13 @@ Additions to existing modules map-concat : map f (concat xss) ≡ concat (map (map f) xss) ``` +* New lemma in `Data.Vec.Relation.Binary.Equality.Cast`: + ```agda + ≈-cong′ : ∀ {f-len : ℕ → ℕ} (f : ∀ {n} → Vec A n → Vec B (f-len n)) + {m n} {xs : Vec A m} {ys : Vec A n} .{eq} → + xs ≈[ eq ] ys → f xs ≈[ _ ] f ys + ``` + * In `Data.Vec.Relation.Binary.Equality.DecPropositional`: ```agda _≡?_ : DecidableEquality (Vec A n) diff --git a/doc/README/Data/Vec/Relation/Binary/Equality/Cast.agda b/doc/README/Data/Vec/Relation/Binary/Equality/Cast.agda index 0852f68166..87379d2339 100644 --- a/doc/README/Data/Vec/Relation/Binary/Equality/Cast.agda +++ b/doc/README/Data/Vec/Relation/Binary/Equality/Cast.agda @@ -20,6 +20,7 @@ open import Data.Nat.Properties open import Data.Vec.Base open import Data.Vec.Properties open import Data.Vec.Relation.Binary.Equality.Cast +open import Function using (_∘_) open import Relation.Binary.PropositionalEquality using (_≡_; refl; sym; cong; module ≡-Reasoning) @@ -187,7 +188,7 @@ example3a-fromList-++-++ {xs = xs} {ys} {zs} eq = begin fromList (xs List.++ ys List.++ zs) ≈⟨ fromList-++ xs ⟩ fromList xs ++ fromList (ys List.++ zs) - ≈⟨ ≈-cong (fromList xs ++_) (cast-++ʳ (List.length-++ ys) (fromList xs)) (fromList-++ ys) ⟩ + ≈⟨ ≈-cong′ (fromList xs ++_) (fromList-++ ys) ⟩ fromList xs ++ fromList ys ++ fromList zs ∎ where open CastReasoning @@ -218,9 +219,7 @@ example4-cong² : ∀ .(eq : (m + 1) + n ≡ n + suc m) a (xs : Vec A m) ys → cast eq (reverse ((xs ++ [ a ]) ++ ys)) ≡ ys ʳ++ reverse (xs ∷ʳ a) example4-cong² {m = m} {n} eq a xs ys = begin reverse ((xs ++ [ a ]) ++ ys) - ≈⟨ ≈-cong reverse (cast-reverse (cong (_+ n) (+-comm 1 m)) ((xs ∷ʳ a) ++ ys)) - (≈-cong (_++ ys) (cast-++ˡ (+-comm 1 m) (xs ∷ʳ a)) - (unfold-∷ʳ-eqFree a xs)) ⟨ + ≈⟨ ≈-cong′ (reverse ∘ (_++ ys)) (unfold-∷ʳ-eqFree a xs) ⟨ reverse ((xs ∷ʳ a) ++ ys) ≈⟨ reverse-++-eqFree (xs ∷ʳ a) ys ⟩ reverse ys ++ reverse (xs ∷ʳ a) @@ -264,7 +263,7 @@ example6a-reverse-∷ʳ {n = n} x xs = begin-≡ reverse (xs ∷ʳ x) ≡⟨ ≈-reflexive refl ⟨ reverse (xs ∷ʳ x) - ≈⟨ ≈-cong reverse (cast-reverse _ _) (unfold-∷ʳ-eqFree x xs) ⟩ + ≈⟨ ≈-cong′ reverse (unfold-∷ʳ-eqFree x xs) ⟩ reverse (xs ++ [ x ]) ≈⟨ reverse-++-eqFree xs [ x ] ⟩ x ∷ reverse xs diff --git a/src/Data/Vec/Properties.agda b/src/Data/Vec/Properties.agda index 23f7caa8b7..225c73c7fa 100644 --- a/src/Data/Vec/Properties.agda +++ b/src/Data/Vec/Properties.agda @@ -375,6 +375,8 @@ lookup∘update′ {i = i} {j} i≢j xs y = lookup∘updateAt′ i j i≢j xs open VecCast public using (cast-is-id; cast-trans) +open VecCast using (≈-cong′) + subst-is-cast : (eq : m ≡ n) (xs : Vec A m) → subst (Vec A) eq xs ≡ cast eq xs subst-is-cast refl xs = sym (cast-is-id refl xs) @@ -398,9 +400,7 @@ map-const (_ ∷ xs) y = cong (y ∷_) (map-const xs y) map-cast : (f : A → B) .(eq : m ≡ n) (xs : Vec A m) → map f (cast eq xs) ≡ cast eq (map f xs) -map-cast {n = zero} f eq [] = refl -map-cast {n = suc _} f eq (x ∷ xs) - = cong (f x ∷_) (map-cast f (suc-injective eq) xs) +map-cast f _ _ = sym (≈-cong′ (map f) refl) map-++ : ∀ (f : A → B) (xs : Vec A m) (ys : Vec A n) → map f (xs ++ ys) ≡ map f xs ++ map f ys @@ -494,13 +494,11 @@ toList-map f (x ∷ xs) = cong (f x List.∷_) (toList-map f xs) cast-++ˡ : ∀ .(eq : m ≡ o) (xs : Vec A m) {ys : Vec A n} → cast (cong (_+ n) eq) (xs ++ ys) ≡ cast eq xs ++ ys -cast-++ˡ {o = zero} eq [] {ys} = cast-is-id refl (cast eq [] ++ ys) -cast-++ˡ {o = suc o} eq (x ∷ xs) {ys} = cong (x ∷_) (cast-++ˡ (cong pred eq) xs) +cast-++ˡ _ _ {ys} = ≈-cong′ (_++ ys) refl cast-++ʳ : ∀ .(eq : n ≡ o) (xs : Vec A m) {ys : Vec A n} → cast (cong (m +_) eq) (xs ++ ys) ≡ xs ++ cast eq ys -cast-++ʳ {m = zero} eq [] {ys} = refl -cast-++ʳ {m = suc m} eq (x ∷ xs) {ys} = cong (x ∷_) (cast-++ʳ eq xs) +cast-++ʳ _ xs = ≈-cong′ (xs ++_) refl lookup-++-< : ∀ (xs : Vec A m) (ys : Vec A n) → ∀ i (i