------------------------------------------------------------------------ -- The Agda standard library -- -- Some Vec-related properties ------------------------------------------------------------------------ module Data.Vec.Properties where open import Algebra open import Category.Applicative.Indexed open import Category.Monad open import Category.Monad.Identity open import Data.Vec open import Data.List.Any using (here; there) renaming (module Membership-≡ to List) open import Data.Nat open import Data.Empty using (⊥-elim) import Data.Nat.Properties as Nat open import Data.Fin as Fin using (Fin; zero; suc; toℕ; fromℕ) open import Data.Fin.Props using (_+′_) open import Data.Empty using (⊥-elim) open import Function open import Function.Inverse using (_↔_) open import Relation.Binary module UsingVectorEquality {s₁ s₂} (S : Setoid s₁ s₂) where private module SS = Setoid S open SS using () renaming (Carrier to A) import Data.Vec.Equality as VecEq open VecEq.Equality S replicate-lemma : ∀ {m} n x (xs : Vec A m) → replicate {n = n} x ++ (x ∷ xs) ≈ replicate {n = 1 + n} x ++ xs replicate-lemma zero x xs = refl (x ∷ xs) replicate-lemma (suc n) x xs = SS.refl ∷-cong replicate-lemma n x xs xs++[]=xs : ∀ {n} (xs : Vec A n) → xs ++ [] ≈ xs xs++[]=xs [] = []-cong xs++[]=xs (x ∷ xs) = SS.refl ∷-cong xs++[]=xs xs map-++-commute : ∀ {b m n} {B : Set b} (f : B → A) (xs : Vec B m) {ys : Vec B n} → map f (xs ++ ys) ≈ map f xs ++ map f ys map-++-commute f [] = refl _ map-++-commute f (x ∷ xs) = SS.refl ∷-cong map-++-commute f xs open import Relation.Binary.PropositionalEquality as P using (_≡_; _≢_; refl; _≗_) open import Relation.Binary.HeterogeneousEquality using (_≅_; refl) -- lookup is an applicative functor morphism. lookup-morphism : ∀ {a n} (i : Fin n) → Morphism (applicative {a = a}) (RawMonad.rawIApplicative IdentityMonad) lookup-morphism i = record { op = lookup i ; op-pure = lookup-replicate i ; op-⊛ = lookup-⊛ i } where lookup-replicate : ∀ {a n} {A : Set a} (i : Fin n) → lookup i ∘ replicate {A = A} ≗ id {A = A} lookup-replicate zero = λ _ → refl lookup-replicate (suc i) = lookup-replicate i lookup-⊛ : ∀ {a b n} {A : Set a} {B : Set b} i (fs : Vec (A → B) n) (xs : Vec A n) → lookup i (fs ⊛ xs) ≡ (lookup i fs $ lookup i xs) lookup-⊛ zero (f ∷ fs) (x ∷ xs) = refl lookup-⊛ (suc i) (f ∷ fs) (x ∷ xs) = lookup-⊛ i fs xs -- tabulate is an inverse of flip lookup. lookup∘tabulate : ∀ {a n} {A : Set a} (f : Fin n → A) (i : Fin n) → lookup i (tabulate f) ≡ f i lookup∘tabulate f zero = refl lookup∘tabulate f (suc i) = lookup∘tabulate (f ∘ suc) i tabulate∘lookup : ∀ {a n} {A : Set a} (xs : Vec A n) → tabulate (flip lookup xs) ≡ xs tabulate∘lookup [] = refl tabulate∘lookup (x ∷ xs) = P.cong (_∷_ x) $ tabulate∘lookup xs -- If you look up an index in allFin n, then you get the index. lookup-allFin : ∀ {n} (i : Fin n) → lookup i (allFin n) ≡ i lookup-allFin = lookup∘tabulate id -- Various lemmas relating lookup and _++_. lookup-++-< : ∀ {a} {A : Set a} {m n} (xs : Vec A m) (ys : Vec A n) i (i