]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/basics/lists/list.ma
1) New file russell with the coercions to activate Russell-style proving.
[helm.git] / matita / matita / lib / basics / lists / list.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/types.ma".
13 include "arithmetics/nat.ma".
14
15 inductive list (A:Type[0]) : Type[0] :=
16   | nil: list A
17   | cons: A -> list A -> list A.
18
19 notation "hvbox(hd break :: tl)"
20   right associative with precedence 47
21   for @{'cons $hd $tl}.
22
23 notation "[ list0 x sep ; ]"
24   non associative with precedence 90
25   for ${fold right @'nil rec acc @{'cons $x $acc}}.
26
27 notation "hvbox(l1 break @ l2)"
28   right associative with precedence 47
29   for @{'append $l1 $l2 }.
30
31 interpretation "nil" 'nil = (nil ?).
32 interpretation "cons" 'cons hd tl = (cons ? hd tl).
33
34 definition not_nil: ∀A:Type[0].list A → Prop ≝
35  λA.λl.match l with [ nil ⇒ True | cons hd tl ⇒ False ].
36
37 theorem nil_cons:
38   ∀A:Type[0].∀l:list A.∀a:A. a::l ≠ [].
39   #A #l #a @nmk #Heq (change with (not_nil ? (a::l))) >Heq //
40 qed.
41
42 (*
43 let rec id_list A (l: list A) on l :=
44   match l with
45   [ nil => []
46   | (cons hd tl) => hd :: id_list A tl ]. *)
47
48 let rec append A (l1: list A) l2 on l1 ≝ 
49   match l1 with
50   [ nil ⇒  l2
51   | cons hd tl ⇒  hd :: append A tl l2 ].
52
53 definition hd ≝ λA.λl: list A.λd:A.
54   match l with [ nil ⇒ d | cons a _ ⇒ a].
55
56 definition tail ≝  λA.λl: list A.
57   match l with [ nil ⇒  [] | cons hd tl ⇒  tl].
58
59 interpretation "append" 'append l1 l2 = (append ? l1 l2).
60
61 theorem append_nil: ∀A.∀l:list A.l @ [] = l.
62 #A #l (elim l) normalize // qed.
63
64 theorem associative_append: 
65  ∀A.associative (list A) (append A).
66 #A #l1 #l2 #l3 (elim l1) normalize // qed.
67
68 (* deleterio per auto 
69 ntheorem cons_append_commute:
70   ∀A:Type.∀l1,l2:list A.∀a:A.
71     a :: (l1 @ l2) = (a :: l1) @ l2.
72 //; nqed. *)
73
74 theorem append_cons:∀A.∀a:A.∀l,l1.l@(a::l1)=(l@[a])@l1.
75 #A #a #l #l1 >associative_append // qed.
76
77 theorem nil_append_elim: ∀A.∀l1,l2: list A.∀P:?→?→Prop. 
78   l1@l2=[] → P (nil A) (nil A) → P l1 l2.
79 #A #l1 #l2 #P (cases l1) normalize //
80 #a #l3 #heq destruct
81 qed.
82
83 theorem nil_to_nil:  ∀A.∀l1,l2:list A.
84   l1@l2 = [] → l1 = [] ∧ l2 = [].
85 #A #l1 #l2 #isnil @(nil_append_elim A l1 l2) /2/
86 qed.
87
88 (**************************** iterators ******************************)
89
90 let rec map (A,B:Type[0]) (f: A → B) (l:list A) on l: list B ≝
91  match l with [ nil ⇒ nil ? | cons x tl ⇒ f x :: (map A B f tl)].
92
93 lemma map_append : ∀A,B,f,l1,l2.
94   (map A B f l1) @ (map A B f l2) = map A B f (l1@l2).
95 #A #B #f #l1 elim l1
96 [ #l2 @refl
97 | #h #t #IH #l2 normalize //
98 ] qed.
99   
100 let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:list A) on l :B ≝  
101  match l with [ nil ⇒ b | cons a l ⇒ f a (foldr A B f b l)].
102  
103 definition filter ≝ 
104   λT.λp:T → bool.
105   foldr T (list T) (λx,l0.if p x then x::l0 else l0) (nil T).
106
107 (* compose f [a1;...;an] [b1;...;bm] = 
108   [f a1 b1; ... ;f an b1; ... ;f a1 bm; f an bm] *)
109  
110 definition compose ≝ λA,B,C.λf:A→B→C.λl1,l2.
111     foldr ?? (λi,acc.(map ?? (f i) l2)@acc) [ ] l1.
112
113 lemma filter_true : ∀A,l,a,p. p a = true → 
114   filter A p (a::l) = a :: filter A p l.
115 #A #l #a #p #pa (elim l) normalize >pa normalize // qed.
116
117 lemma filter_false : ∀A,l,a,p. p a = false → 
118   filter A p (a::l) = filter A p l.
119 #A #l #a #p #pa (elim l) normalize >pa normalize // qed.
120
121 theorem eq_map : ∀A,B,f,g,l. (∀x.f x = g x) → map A B f l = map A B g l.
122 #A #B #f #g #l #eqfg (elim l) normalize // qed.
123
124 let rec dprodl (A:Type[0]) (f:A→Type[0]) (l1:list A) (g:(∀a:A.list (f a))) on l1 ≝
125 match l1 with
126   [ nil ⇒ nil ?
127   | cons a tl ⇒ (map ??(mk_Sig ?? a) (g a)) @ dprodl A f tl g
128   ].
129
130 (**************************** length ******************************)
131
132 let rec length (A:Type[0]) (l:list A) on l ≝ 
133   match l with 
134     [ nil ⇒ 0
135     | cons a tl ⇒ S (length A tl)].
136
137 notation "|M|" non associative with precedence 60 for @{'norm $M}.
138 interpretation "norm" 'norm l = (length ? l).
139
140 lemma length_append: ∀A.∀l1,l2:list A. 
141   |l1@l2| = |l1|+|l2|.
142 #A #l1 elim l1 // normalize /2/
143 qed.
144
145 (****************************** nth ********************************)
146 let rec nth n (A:Type[0]) (l:list A) (d:A)  ≝  
147   match n with
148     [O ⇒ hd A l d
149     |S m ⇒ nth m A (tail A l) d].
150
151 lemma nth_nil: ∀A,a,i. nth i A ([]) a = a.
152 #A #a #i elim i normalize //
153 qed.
154
155 (****************************** nth_opt ********************************)
156 let rec nth_opt (A:Type[0]) (n:nat) (l:list A) on l : option A ≝
157 match l with
158 [ nil ⇒ None ?
159 | cons h t ⇒ match n with [ O ⇒ Some ? h | S m ⇒ nth_opt A m t ]
160 ].
161
162 (**************************** All *******************************)
163
164 let rec All (A:Type[0]) (P:A → Prop) (l:list A) on l : Prop ≝
165 match l with
166 [ nil ⇒ True
167 | cons h t ⇒ P h ∧ All A P t
168 ].
169
170 lemma All_mp : ∀A,P,Q. (∀a.P a → Q a) → ∀l. All A P l → All A Q l.
171 #A #P #Q #H #l elim l normalize //
172 #h #t #IH * /3/
173 qed.
174
175 lemma All_nth : ∀A,P,n,l.
176   All A P l →
177   ∀a.
178   nth_opt A n l = Some A a →
179   P a.
180 #A #P #n elim n
181 [ * [ #_ #a #E whd in E:(??%?); destruct
182     | #hd #tl * #H #_ #a #E whd in E:(??%?); destruct @H
183     ]
184 | #m #IH *
185   [ #_ #a #E whd in E:(??%?); destruct
186   | #hd #tl * #_ whd in ⊢ (? → ∀_.??%? → ?); @IH
187   ]
188 ] qed.
189
190 (**************************** Exists *******************************)
191
192 let rec Exists (A:Type[0]) (P:A → Prop) (l:list A) on l : Prop ≝
193 match l with
194 [ nil ⇒ False
195 | cons h t ⇒ (P h) ∨ (Exists A P t)
196 ].
197
198 lemma Exists_append : ∀A,P,l1,l2.
199   Exists A P (l1 @ l2) → Exists A P l1 ∨ Exists A P l2.
200 #A #P #l1 elim l1
201 [ normalize /2/
202 | #h #t #IH #l2 *
203   [ #H /3/
204   | #H cases (IH l2 H) /3/
205   ]
206 ] qed.
207
208 lemma Exists_append_l : ∀A,P,l1,l2.
209   Exists A P l1 → Exists A P (l1@l2).
210 #A #P #l1 #l2 elim l1
211 [ *
212 | #h #t #IH *
213   [ #H %1 @H
214   | #H %2 @IH @H
215   ]
216 ] qed.
217
218 lemma Exists_append_r : ∀A,P,l1,l2.
219   Exists A P l2 → Exists A P (l1@l2).
220 #A #P #l1 #l2 elim l1
221 [ #H @H
222 | #h #t #IH #H %2 @IH @H
223 ] qed.
224
225 lemma Exists_add : ∀A,P,l1,x,l2. Exists A P (l1@l2) → Exists A P (l1@x::l2).
226 #A #P #l1 #x #l2 elim l1
227 [ normalize #H %2 @H
228 | #h #t #IH normalize * [ #H %1 @H | #H %2 @IH @H ]
229 qed.
230
231 lemma Exists_mid : ∀A,P,l1,x,l2. P x → Exists A P (l1@x::l2).
232 #A #P #l1 #x #l2 #H elim l1
233 [ %1 @H
234 | #h #t #IH %2 @IH
235 ] qed.
236
237 lemma Exists_map : ∀A,B,P,Q,f,l.
238 Exists A P l →
239 (∀a.P a → Q (f a)) →
240 Exists B Q (map A B f l).
241 #A #B #P #Q #f #l elim l //
242 #h #t #IH * [ #H #F %1 @F @H | #H #F %2 @IH [ @H | @F ] ] qed.
243
244 lemma Exists_All : ∀A,P,Q,l.
245   Exists A P l →
246   All A Q l →
247   ∃x. P x ∧ Q x.
248 #A #P #Q #l elim l [ * | #hd #tl #IH * [ #H1 * #H2 #_ %{hd} /2/ | #H1 * #_ #H2 @IH // ]
249 qed.
250
251 (**************************** fold *******************************)
252
253 let rec fold (A,B:Type[0]) (op:B → B → B) (b:B) (p:A→bool) (f:A→B) (l:list A) on l :B ≝  
254  match l with 
255   [ nil ⇒ b 
256   | cons a l ⇒
257      if p a then op (f a) (fold A B op b p f l)
258      else fold A B op b p f l].
259       
260 notation "\fold  [ op , nil ]_{ ident i ∈ l | p} f"
261   with precedence 80
262 for @{'fold $op $nil (λ${ident i}. $p) (λ${ident i}. $f) $l}.
263
264 notation "\fold [ op , nil ]_{ident i ∈ l } f"
265   with precedence 80
266 for @{'fold $op $nil (λ${ident i}.true) (λ${ident i}. $f) $l}.
267
268 interpretation "\fold" 'fold op nil p f l = (fold ? ? op nil p f l).
269
270 theorem fold_true: 
271 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A→B. p a = true → 
272   \fold[op,nil]_{i ∈ a::l| p i} (f i) = 
273     op (f a) \fold[op,nil]_{i ∈ l| p i} (f i). 
274 #A #B #a #l #p #op #nil #f #pa normalize >pa // qed.
275
276 theorem fold_false: 
277 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f.
278 p a = false → \fold[op,nil]_{i ∈ a::l| p i} (f i) = 
279   \fold[op,nil]_{i ∈ l| p i} (f i).
280 #A #B #a #l #p #op #nil #f #pa normalize >pa // qed.
281
282 theorem fold_filter: 
283 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A →B.
284   \fold[op,nil]_{i ∈ l| p i} (f i) = 
285     \fold[op,nil]_{i ∈ (filter A p l)} (f i).
286 #A #B #a #l #p #op #nil #f elim l //  
287 #a #tl #Hind cases(true_or_false (p a)) #pa 
288   [ >filter_true // > fold_true // >fold_true //
289   | >filter_false // >fold_false // ]
290 qed.
291
292 record Aop (A:Type[0]) (nil:A) : Type[0] ≝
293   {op :2> A → A → A; 
294    nill:∀a. op nil a = a; 
295    nilr:∀a. op a nil = a;
296    assoc: ∀a,b,c.op a (op b c) = op (op a b) c
297   }.
298
299 theorem fold_sum: ∀A,B. ∀I,J:list A.∀nil.∀op:Aop B nil.∀f.
300   op (\fold[op,nil]_{i∈I} (f i)) (\fold[op,nil]_{i∈J} (f i)) =
301     \fold[op,nil]_{i∈(I@J)} (f i).
302 #A #B #I #J #nil #op #f (elim I) normalize 
303   [>nill //|#a #tl #Hind <assoc //]
304 qed.
305
306 (********************** lhd and ltl ******************************)
307
308 let rec lhd (A:Type[0]) (l:list A) n on n ≝ match n with
309    [ O   ⇒ nil …
310    | S n ⇒ match l with [ nil ⇒ nil … | cons a l ⇒ a :: lhd A l n ]
311    ].
312
313 let rec ltl (A:Type[0]) (l:list A) n on n ≝ match n with
314    [ O   ⇒ l
315    | S n ⇒ ltl A (tail … l) n
316    ].
317
318 lemma lhd_nil: ∀A,n. lhd A ([]) n = [].
319 #A #n elim n //
320 qed.
321
322 lemma ltl_nil: ∀A,n. ltl A ([]) n = [].
323 #A #n elim n normalize //
324 qed.
325
326 lemma lhd_cons_ltl: ∀A,n,l. lhd A l n @ ltl A l n = l.
327 #A #n elim n -n //
328 #n #IHn #l elim l normalize //
329 qed.
330
331 lemma length_ltl: ∀A,n,l. |ltl A l n| = |l| - n.
332 #A #n elim n -n /2/
333 #n #IHn *; normalize /2/
334 qed.
335
336 (********************** find ******************************)
337 let rec find (A,B:Type[0]) (f:A → option B) (l:list A) on l : option B ≝
338 match l with
339 [ nil ⇒ None B
340 | cons h t ⇒
341     match f h with
342     [ None ⇒ find A B f t
343     | Some b ⇒ Some B b
344     ]
345 ].
346
347 (********************** position_of ******************************)
348 let rec position_of_aux (A:Type[0]) (found: A → bool) (l:list A) (acc:nat) on l : option nat ≝
349 match l with
350 [ nil ⇒ None ?
351 | cons h t ⇒
352    match found h with [true ⇒ Some … acc | false ⇒ position_of_aux … found t (S acc)]].
353
354 definition position_of: ∀A:Type[0]. (A → bool) → list A → option nat ≝
355  λA,found,l. position_of_aux A found l 0.
356
357
358 (********************** make_list ******************************)
359 let rec make_list (A:Type[0]) (a:A) (n:nat) on n : list A ≝
360 match n with
361 [ O ⇒ [ ]
362 | S m ⇒ a::(make_list A a m)
363 ].