Library Coq.Vectors.VectorSpec


N.B.: Using this encoding of vectors is discouraged. See <https://github.com/coq/coq/blob/master/theories/Vectors/Vector.v>.

Proofs of specification for functions defined over Vector
Initial Author: Pierre Boutillier Institution: PPS, INRIA 12/2010

#[local] Set Warnings "-stdlib-vector".
Require Fin List.
Require Import VectorDef PeanoNat Eqdep_dec.
Import VectorNotations EqNotations.

Definition cons_inj {A} {a1 a2} {n} {v1 v2 : t A n}
 (eq : a1 :: v1 = a2 :: v2) : a1 = a2 /\ v1 = v2 :=
   match eq in _ = x return caseS _ (fun a2' _ v2' => fun v1' => a1 = a2' /\ v1' = v2') x v1
   with | eq_refl => conj eq_refl eq_refl
   end.

Lemma nil_spec {A} (v : t A 0) : v = [].

Lemma eta {A} {n} (v : t A (S n)) : v = hd v :: tl v.

Lemmas are done for functions that use Fin.t but thanks to Peano_dec.le_unique, all is true for the one that use lt

Properties of nth and nth_order


Lemma eq_nth_iff A n (v1 v2: t A n):
  (forall p1 p2, p1 = p2 -> v1 [@ p1 ] = v2 [@ p2 ]) <-> v1 = v2.

Lemma nth_order_hd A: forall n (v : t A (S n)) (H : 0 < S n),
  nth_order v H = hd v.

Lemma nth_order_tl A: forall n k (v : t A (S n)) (H : k < n) (HS : S k < S n),
  nth_order (tl v) H = nth_order v HS.

Lemma nth_order_last A: forall n (v: t A (S n)) (H: n < S n),
  nth_order v H = last v.

Lemma nth_order_ext A: forall n k (v : t A n) (H1 H2 : k < n),
  nth_order v H1 = nth_order v H2.

Lemma nth_append_L A: forall n m (v : t A n) (w : t A m) p,
  (v ++ w) [@ Fin.L m p] = v [@ p].

Lemma nth_append_R A: forall n m (v : t A n) (w : t A m) p,
  (v ++ w) [@ Fin.R n p] = w [@ p].

Lemma In_nth A: forall n (v : t A n) p,
  In (nth v p) v.

Properties of replace


Lemma nth_replace_eq A: forall n p (v : t A n) a,
  nth (replace v p a) p = a.

Lemma nth_replace_neq A: forall n p q, p <> q -> forall (v : t A n) a,
  nth (replace v q a) p = nth v p.

Lemma nth_order_replace_eq A: forall n k (v : t A n) a (H1 : k < n) (H2 : k < n),
  nth_order (replace v (Fin.of_nat_lt H2) a) H1 = a.

Lemma nth_order_replace_neq A: forall n k1 k2, k1 <> k2 ->
  forall (v : t A n) a (H1 : k1 < n) (H2 : k2 < n),
  nth_order (replace v (Fin.of_nat_lt H2) a) H1 = nth_order v H1.

Lemma replace_id A: forall n p (v : t A n),
  replace v p (nth v p) = v.

Lemma replace_replace_eq A: forall n p (v : t A n) a b,
  replace (replace v p a) p b = replace v p b.

Lemma replace_replace_neq A: forall n p1 p2 (v : t A n) a b, p1 <> p2 ->
  replace (replace v p1 a) p2 b = replace (replace v p2 b) p1 a.

Lemma replace_append_L A: forall n m (v : t A n) (w : t A m) p a,
  replace (v ++ w) (Fin.L m p) a = (replace v p a) ++ w.

Lemma replace_append_R A: forall n m (v : t A n) (w : t A m) p a,
  replace (v ++ w) (Fin.R n p) a = v ++ (replace w p a).

Properties of const


Lemma const_nth A (a: A) n (p: Fin.t n): (const a n)[@ p] = a.

Lemma append_const A (a : A) n m : (const a n) ++ (const a m) = const a (n + m).

Properties of map


Lemma map_id A: forall n (v : t A n),
  map (fun x => x) v = v.

Lemma map_map A B C: forall (f:A->B) (g:B->C) n (v : t A n),
  map g (map f v) = map (fun x => g (f x)) v.

Lemma map_ext_in A B: forall (f g:A->B) n (v : t A n),
  (forall a, In a v -> f a = g a) -> map f v = map g v.

Lemma map_ext A B: forall (f g:A->B), (forall a, f a = g a) ->
  forall n (v : t A n), map f v = map g v.

Lemma nth_map {A B} (f: A -> B) {n} v (p1 p2: Fin.t n) (eq: p1 = p2):
  (map f v) [@ p1] = f (v [@ p2]).

Lemma map_append A B: forall (f:A->B) n m (v: t A n) (w: t A m),
  (map f (v ++ w)) = map f v ++ map f w.

Lemma nth_map2 {A B C} (f: A -> B -> C) {n} v w (p1 p2 p3: Fin.t n):
  p1 = p2 -> p2 = p3 -> (map2 f v w) [@p1] = f (v[@p2]) (w[@p3]).

Lemma map2_ext A B C: forall (f g:A->B->C), (forall a b, f a b = g a b) ->
  forall n (v : t A n) (w : t B n), map2 f v w = map2 g v w.

Properties of fold_left


Lemma fold_left_right_assoc_eq {A B} {f: A -> B -> A}
  (assoc: forall a b c, f (f a b) c = f (f a c) b)
  {n} (v: t B n): forall a, fold_left f a v = fold_right (fun x y => f y x) v a.

Properties of take


Lemma take_O : forall {A} {n} le (v:t A n), take 0 le v = [].

Lemma take_idem : forall {A} p n (v:t A n) le le',
  take p le' (take p le v) = take p le v.

Lemma take_app : forall {A} {n} (v:t A n) {m} (w:t A m) le, take n le (append v w) = v.

Lemma take_prf_irr : forall {A} p {n} (v:t A n) le le', take p le v = take p le' v.

Properties of uncons and splitat


Lemma uncons_cons {A} : forall {n : nat} (a : A) (v : t A n),
  uncons (a::v) = (a,v).


Lemma append_comm_cons {A} : forall {n m : nat} (v : t A n) (w : t A m) (a : A),
  a :: (v ++ w) = (a :: v) ++ w.

Lemma splitat_append {A} : forall {n m : nat} (v : t A n) (w : t A m),
  splitat n (v ++ w) = (v, w).

Lemma append_splitat {A} : forall {n m : nat} (v : t A n) (w : t A m) (vw : t A (n+m)),
  splitat n vw = (v, w) -> vw = v ++ w.

Lemma append_inj {A} : forall {n m : nat} (v v' : t A n) (w w' : t A m),
  v ++ w = v' ++ w' -> v = v' /\ w = w'.

Properties of In


Lemma In_cons_iff A: forall n a b (v : t A n),
  In a (b :: v) <-> (b = a \/ In a v).

Properties of Forall and Forall2


Lemma Forall_impl A: forall (P Q : A -> Prop), (forall a, P a -> Q a) ->
  forall n (v : t A n), Forall P v -> Forall Q v.

Lemma Forall_forall A: forall P n (v : t A n),
  Forall P v <-> forall a, In a v -> P a.

Lemma Forall_cons_iff A: forall (P : A -> Prop) a n (v : t A n),
  Forall P (a :: v) <-> P a /\ Forall P v.

Lemma Forall_map A B: forall (P : B -> Prop) (f : A -> B) n (v : t A n),
  Forall P (map f v) <-> Forall (fun x => P (f x)) v.

Lemma Forall_append A: forall P n m (v : t A n) (w : t A m),
  Forall P v -> Forall P w -> Forall P (append v w).

Lemma Forall_nth A: forall P n (v : t A n),
  Forall P v <-> forall p, P (nth v p).

Lemma Forall_nth_order A: forall P n (v : t A n),
  Forall P v <-> forall i (Hi : i < n), P (nth_order v Hi).

Lemma Forall2_nth A: forall P n (v1 v2 : t A n),
  Forall2 P v1 v2 <-> forall p, P (nth v1 p) (nth v2 p).

Lemma Forall2_nth_order A: forall P n (v1 v2 : t A n),
     Forall2 P v1 v2
 <-> forall i (Hi1 : i < n) (Hi2 : i < n), P (nth_order v1 Hi1) (nth_order v2 Hi2).

Lemma Forall2_append A B: forall P n m (v : t A n) (v' : t B n) (w : t A m) (w' : t B m),
  Forall2 P v v' -> Forall2 P w w' -> Forall2 P (append v w) (append v' w').

Properties of shiftin and shiftrepeat


Lemma shiftin_nth A a n (v: t A n) k1 k2 (eq: k1 = k2):
  nth (shiftin a v) (Fin.L_R 1 k1) = nth v k2.

Lemma shiftin_last A a n (v: t A n): last (shiftin a v) = a.

Lemma shiftrepeat_nth A: forall n k (v: t A (S n)),
  nth (shiftrepeat v) (Fin.L_R 1 k) = nth v k.

Lemma shiftrepeat_last A: forall n (v: t A (S n)), last (shiftrepeat v) = last v.

Lemma map_shiftin A B: forall (f : A -> B) a n (v : t A n),
  map f (shiftin a v) = shiftin (f a) (map f v).

Lemma fold_right_shiftin A B: forall (f : A -> B -> B) a b n (v : t A n),
  fold_right f (shiftin b v) a = fold_right f v (f b a).

Lemma In_shiftin A: forall a b n (v : t A n),
  In a (shiftin b v) <-> (b = a \/ In a v).

Lemma Forall_shiftin A: forall (P : A -> Prop) a n (v : t A n),
  Forall P (shiftin a v) <-> (P a /\ Forall P v).

Properties of rev


Lemma rev_nil A: rev (nil A) = [].

Lemma rev_cons A: forall a n (v : t A n), rev (a :: v) = shiftin a (rev v).

Lemma rev_shiftin A: forall a n (v : t A n),
  rev (shiftin a v) = a :: rev v.

Lemma rev_rev A: forall n (v : t A n), rev (rev v) = v.

Lemma map_rev A B: forall (f : A -> B) n (v : t A n),
  map f (rev v) = rev (map f v).

Lemma fold_left_rev_right A B: forall (f:A->B->B) n (v : t A n) a,
  fold_right f (rev v) a = fold_left (fun x y => f y x) a v.

Lemma In_rev A: forall a n (v : t A n),
  In a (rev v) <-> In a v.

Lemma Forall_rev A: forall (P : A -> Prop) n (v : t A n),
  Forall P (rev v) <-> Forall P v.

Properties of to_list


Lemma to_list_of_list_opp {A} (l: list A): to_list (of_list l) = l.

Lemma length_to_list A n (v : t A n): length (to_list v) = n.

Lemma of_list_to_list_opp A n (v: t A n):
  rew length_to_list _ _ _ in of_list (to_list v) = v.

Lemma to_list_nil A : to_list (nil A) = List.nil.

Lemma to_list_cons A h n (v : t A n):
  to_list (cons A h n v) = List.cons h (to_list v).

Lemma to_list_nil_iff A v:
  to_list v = List.nil <-> v = nil A.

Lemma to_list_inj A n (v w : t A n):
  to_list v = to_list w -> v = w.

Lemma to_list_hd A n (v : t A (S n)) d:
  hd v = List.hd d (to_list v).

Lemma to_list_last A n (v : t A (S n)) d:
  last v = List.last (to_list v) d.

Lemma to_list_const A (a : A) n:
  to_list (const a n) = List.repeat a n.

Lemma to_list_nth_order A n (v : t A n) p (H : p < n) d:
  nth_order v H = List.nth p (to_list v) d.

Lemma to_list_tl A n (v : t A (S n)):
  to_list (tl v) = List.tl (to_list v).

Lemma to_list_append A n m (v1 : t A n) (v2 : t A m):
  to_list (append v1 v2) = List.app (to_list v1) (to_list v2).

Lemma to_list_rev_append_tail A n m (v1 : t A n) (v2 : t A m):
  to_list (rev_append_tail v1 v2) = List.rev_append (to_list v1) (to_list v2).

Lemma to_list_rev_append A n m (v1 : t A n) (v2 : t A m):
  to_list (rev_append v1 v2) = List.rev_append (to_list v1) (to_list v2).

Lemma to_list_rev A n (v : t A n):
  to_list (rev v) = List.rev (to_list v).

Lemma to_list_map A B (f : A -> B) n (v : t A n) :
  to_list (map f v) = List.map f (to_list v).

Lemma to_list_fold_left A B f (b : B) n (v : t A n):
  fold_left f b v = List.fold_left f (to_list v) b.

Lemma to_list_fold_right A B f (b : B) n (v : t A n):
  fold_right f v b = List.fold_right f b (to_list v).

Lemma to_list_Forall A P n (v : t A n):
  Forall P v <-> List.Forall P (to_list v).

Lemma to_list_Exists A P n (v : t A n):
  Exists P v <-> List.Exists P (to_list v).

Lemma to_list_In A a n (v : t A n):
  In a v <-> List.In a (to_list v).

Lemma to_list_Forall2 A B P n (v1 : t A n) (v2 : t B n) :
  Forall2 P v1 v2 <-> List.Forall2 P (to_list v1) (to_list v2).