]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/basics/vectors.ma
e680e91a7170a48ce231d9c4408f7028f2bbe2e8
[helm.git] / matita / matita / lib / basics / vectors.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic   
3     ||A||  Library of Mathematics, developed at the Computer Science 
4     ||T||  Department of the University of Bologna, Italy.           
5     ||I||                                                            
6     ||T||  
7     ||A||  
8     \   /  This file is distributed under the terms of the       
9      \ /   GNU General Public License Version 2   
10       V_____________________________________________________________*)
11
12 include "basics/finset.ma".
13
14 record Vector (A:Type[0]) (n:nat): Type[0] ≝ 
15 { vec :> list A;
16   len: length ? vec = n
17 }.
18
19 lemma Vector_eq : ∀A,n,v1,v2.
20   vec A n v1 = vec A n v2 → v1 = v2.
21 #A #n * #l1 #H1 * #l2 #H2 #eql1l2 generalize in match H1; 
22 -H1 >eql1l2 //
23 qed.
24
25 definition vec_tail ≝ λA.λn.λv:Vector A n.
26 mk_Vector A (pred n) (tail A v) ?.
27 >length_tail >(len A n v) //
28 qed.
29
30 definition vec_cons ≝ λA.λa.λn.λv:Vector A n.
31 mk_Vector A (S n) (cons A a v) ?.
32 normalize >(len A n v) //
33 qed.
34
35 definition vec_hd ≝ λA.λn.λv:Vector A (S n).
36 hd A v ?. elim v * [normalize #H destruct | #a #H #eq @a] 
37 qed.
38
39 lemma vec_expand: ∀A,n,v. 
40  v = vec_cons A (vec_hd A n v) n (vec_tail A (S n) v).
41 #A #n * #l cases l [normalize in ⊢ (%→?); #H destruct  |//]
42 qed.
43
44 lemma vector_nil: ∀A.∀v:Vector A 0. 
45   v = mk_Vector A 0 (nil A) (refl ??).
46 #A * * // #a #tl normalize #H destruct
47 qed. 
48
49 definition vec_append ≝ λA.λn1,n2.λv1:Vector A n1.λv2: Vector A n2.
50 mk_Vector A (n1+n2) (v1@v2).
51
52 definition vec_map ≝ λA,B.λf:A→B.λn.λv:Vector A n.
53 mk_Vector B n (map ?? f v) 
54   (trans_eq … (length_map …) (len A n v)).
55   
56 lemma nth_default : ∀A,i,n.∀v:Vector A n.∀d1,d2. i < n →
57   nth i ? v d1 = nth i ? v d2.
58 #A #i elim i -i
59   [#n #v #d1 #d2 #ltOn lapply v @(lt_O_n_elim … ltOn)
60    -v #m #v >(vec_expand … v) //
61   |#i #Hind #n #v #d1 #d2 #ltn lapply ltn lapply v @(lt_O_n_elim … (ltn_to_ltO … ltn))
62    -v -ltn #m #v #ltim >(vec_expand … v) @(Hind m (vec_tail A (S m) v) d1 d2 ?) 
63    @le_S_S_to_le //
64   ]
65 qed.
66   
67 lemma eq_vec: ∀A,n.∀v1,v2:Vector A n.∀d. 
68   (∀i. i < n → nth i A v1 d = nth i A v2 d) → v1 = v2.
69 #A #n elim n -n 
70   [#v1 #v2 #H >(vector_nil A v1) >(vector_nil A v2) //
71   |#n #Hind #v1 #v2 #d #H >(vec_expand … v1) >(vec_expand … v2)
72    >(Hind (vec_tail … v1) (vec_tail … v2) d)
73     [cut (vec_hd A n v1 = vec_hd A n v2) //
74      cut (∀i,d1,d2. i < S n → nth i A v1 d1 = nth i A v2 d2)
75        [#i #d1 #d2 #Hi >(nth_default ????? d) // >(nth_default ???? d2 d) // @H //]
76      -H #H @(H 0) //  
77     |#i #ltin @(H (S i)) @le_S_S //
78     ]
79   ]
80 qed.
81
82 lemma nth_vec_map :
83   ∀A,B,f,i,n.∀v:Vector A n.∀d.
84   f (nth i ? v d) = nth i ? (vec_map A B f n v) (f d).
85 #A #B #f #i elim i
86 [ *
87   [ #v #d >(vector_nil … v) %
88   | #n0 #v #d >(vec_expand … v) % ]
89 | #i0 #IH *
90   [ #v #d >(vector_nil … v) normalize cases i0 // 
91   | #n #v #d >(vec_expand … v) whd in ⊢ (??(?%)%);
92     >(IH n (vec_tail A (S n) v) d) % ] ]
93 qed.
94
95
96 (* mapi: map with index to move in list.ma *)
97 let rec change_vec (A:Type[0]) (n:nat) on n ≝
98 match n return λn0.∀v:Vector A n0.A→nat→Vector A n0 with
99 [ O ⇒ λv,a,i.v
100 | S m ⇒ λv,a,i.
101   match i with
102   [ O ⇒ vec_cons A a m (vec_tail … v) 
103   | S j ⇒ vec_cons A (vec_hd A m v) m (change_vec A m (vec_tail … v) a j)
104   ]
105 ].
106
107 lemma nth_change_vec : ∀A,i,n,v,a,d. i < n →
108   nth i ? (change_vec A n v a i) d = a.
109 #A #i elim i
110   [#n #v #a #d #ltOn lapply v -v @(lt_O_n_elim n ltOn ??) //
111   |#m #Hind #n #v #a #d #Hlt 
112    lapply Hlt lapply v @(lt_O_n_elim … (ltn_to_ltO … Hlt)) 
113    #p #v #ltmp @Hind @le_S_S_to_le // 
114   ]
115 qed.
116
117 lemma nth_change_vec_neq : ∀A,j,i,n,v,a,d. i ≠ j →
118   nth j ? (change_vec A n v a i) d = nth j ? v d.
119 #A #j elim j
120   [#i * // #n #v #a #d cases i
121     [#H @False_ind @(absurd ?? H) //
122     |#i0 #_ >(vec_expand ?? v) in ⊢ (???%); //
123     ] 
124   |#m #Hind #i * // cases i // #i0 #n #v #a #d #neqim 
125    whd in ⊢(??%?); whd in match (tail ??); >Hind
126     [>(vec_expand ??v) in ⊢ (???%); // |@(not_to_not … neqim) // ]
127   ]
128 qed.
129
130 lemma change_vec_same : ∀sig,n,v,i,d.
131   change_vec sig n v (nth i ? v d) i = v.
132 #sig #n #v #i #d @(eq_vec … d)
133 #i0 #Hi0 cases (decidable_eq_nat i i0) #Hi0
134 [ >Hi0 >nth_change_vec //
135 | >nth_change_vec_neq //
136 ]
137 qed.
138
139 lemma change_vec_cons_tail :∀A,n,vA,a,b,i.
140   change_vec A (S n) (vec_cons ? a n vA) b (S i) =
141   vec_cons ? a n (change_vec A n vA b i).
142 #A #n #vA cases vA //
143 qed.
144
145 lemma change_vec_commute : ∀A,n,v,a,b,i,j. i ≠ j → 
146   change_vec A n (change_vec A n v a i) b j
147   = change_vec A n (change_vec A n v b j) a i.
148 #A #n #v #a #b #i #j #Hij @(eq_vec … a)
149 #k #Hk cases (decidable_eq_nat k i) #Hki
150 [ >Hki >nth_change_vec // >(nth_change_vec_neq ??????? (sym_not_eq … Hij))
151   >nth_change_vec //
152 | cases (decidable_eq_nat k j) #Hkj
153   [ >Hkj >nth_change_vec // >(nth_change_vec_neq ??????? Hij) >nth_change_vec //
154   | >(nth_change_vec_neq ??????? (sym_not_eq … Hki))
155     >(nth_change_vec_neq ??????? (sym_not_eq … Hkj))
156     >(nth_change_vec_neq ??????? (sym_not_eq … Hki))
157     >(nth_change_vec_neq ??????? (sym_not_eq … Hkj)) //
158   ]
159 ]
160 qed.
161
162 lemma change_vec_change_vec : ∀A,n,v,a,b,i. 
163   change_vec A n (change_vec A n v a i) b i = change_vec A n v b i.
164 #A #n #v #a #b #i @(eq_vec … a) #i0 #Hi0
165 cases (decidable_eq_nat i i0) #Hii0
166 [ >Hii0 >nth_change_vec // >nth_change_vec //
167 | >nth_change_vec_neq // >nth_change_vec_neq //
168   >nth_change_vec_neq // ]
169 qed.
170
171
172 (*
173 lemma length_make_listi: ∀A,a,n,i. 
174   |make_listi A a n i| = n.
175 #A #a #n elim n // #m #Hind normalize //
176 qed.
177 definition change_vec ≝ λA,n,v,a,i.
178   make_veci A (λj.if (eqb i j) then a else (nth j A v a)) n 0.
179
180 let rec mapi (A,B:Type[0]) (f: nat → A → B) (l:list A) (i:nat) on l: list B ≝
181  match l with 
182  [ nil ⇒ nil ? 
183  | cons x tl ⇒ f i x :: (mapi A B f tl (S i))].
184  
185 lemma length_mapi: ∀A,B,l.∀f:nat→A→B.∀i. 
186   |mapi ?? f l i| = |l|.
187 #A #B #l #f elim l // #a #tl #Hind normalize //
188 qed.
189
190 let rec make_listi (A:Type[0]) (a:nat→A) (n,i:nat) on n : list A ≝
191 match n with
192 [ O ⇒ [ ]
193 | S m ⇒ a i::(make_listi A a m (S i))
194 ].
195
196 lemma length_make_listi: ∀A,a,n,i. 
197   |make_listi A a n i| = n.
198 #A #a #n elim n // #m #Hind normalize //
199 qed.
200
201 definition vec_mapi ≝ λA,B.λf:nat→A→B.λn.λv:Vector A n.λi.
202 mk_Vector B n (mapi ?? f v i) 
203   (trans_eq … (length_mapi …) (len A n v)).
204   
205 definition make_veci ≝ λA.λa:nat→A.λn,i.
206 mk_Vector A n (make_listi A a n i) (length_make_listi A a n i).
207 *)
208
209 let rec pmap A B C (f:A→B→C) l1 l2 on l1 ≝
210    match l1 with
211    [ nil ⇒ nil C
212    | cons a tlA ⇒ 
213      match l2 with
214      [nil ⇒ nil C
215      |cons b tlB ⇒ (f a b)::pmap A B C f tlA tlB
216      ]
217    ].
218
219 lemma length_pmap: ∀A,B,C.∀f:A→B→C.∀l1,l2.
220 length C (pmap  A B C f l1 l2) = 
221   min (length A l1) (length B l2).
222 #A #B #C #f #l1 elim l1 // #a #tlA #Hind #l2 cases l2 //
223 #b #tlB lapply (Hind tlB) normalize 
224 cases (true_or_false (leb (length A tlA) (length B tlB))) #H >H
225 normalize //
226 qed.
227
228 definition pmap_vec ≝ λA,B,C.λf:A→B→C.λn.λvA:Vector A n.λvB:Vector B n.
229 mk_Vector C n (pmap A B C f vA vB) ?.
230 >length_pmap >(len A n vA) >(len B n vB) normalize 
231 >(le_to_leb_true … (le_n n)) //
232 qed.
233
234 lemma pmap_vec_cons :∀A,B,C,f,n,vA,vB,a,b.
235   pmap_vec A B C f (S n) (vec_cons ? a n vA) (vec_cons ? b n vB) =
236   vec_cons ? (f a b) n (pmap_vec A B C f n vA vB).
237 // qed.
238
239 lemma pmap_change : ∀A,B,C,f,n,vA,vB,a,b,i.
240   pmap_vec A B C f n (change_vec ? n vA a i) (change_vec ? n vB b i) =
241   change_vec ? n (pmap_vec A B C f n vA vB) (f a b) i.
242 #A #B #C #f #n elim n //
243 #m #Hind #vA #vB #a #b >(vec_expand … vA) >(vec_expand … vB) * //
244 #i >pmap_vec_cons >pmap_vec_cons >change_vec_cons_tail @eq_f @Hind 
245 qed.