]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/basics/lists/list.ma
a4ba4cce1ee217651b21a3779204aa3637185af0
[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 term 19 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 is_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 (is_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 definition option_hd ≝ 
60   λA.λl:list A. match l with
61   [ nil ⇒ None ?
62   | cons a _ ⇒ Some ? a ].
63
64 interpretation "append" 'append l1 l2 = (append ? l1 l2).
65
66 theorem append_nil: ∀A.∀l:list A.l @ [] = l.
67 #A #l (elim l) normalize // qed.
68
69 theorem associative_append: 
70  ∀A.associative (list A) (append A).
71 #A #l1 #l2 #l3 (elim l1) normalize // qed.
72
73 theorem append_cons:∀A.∀a:A.∀l,l1.l@(a::l1)=(l@[a])@l1.
74 #A #a #l #l1 >associative_append // qed.
75
76 theorem nil_append_elim: ∀A.∀l1,l2: list A.∀P:?→?→Prop. 
77   l1@l2=[] → P (nil A) (nil A) → P l1 l2.
78 #A #l1 #l2 #P (cases l1) normalize //
79 #a #l3 #heq destruct
80 qed.
81
82 theorem nil_to_nil:  ∀A.∀l1,l2:list A.
83   l1@l2 = [] → l1 = [] ∧ l2 = [].
84 #A #l1 #l2 #isnil @(nil_append_elim A l1 l2) /2/
85 qed.
86
87 lemma cons_injective_l : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → a1 = a2.
88 #A #a1 #a2 #l1 #l2 #Heq destruct //
89 qed.
90
91 lemma cons_injective_r : ∀A.∀a1,a2:A.∀l1,l2.a1::l1 = a2::l2 → l1 = l2.
92 #A #a1 #a2 #l1 #l2 #Heq destruct //
93 qed.
94
95 (* option cons *)
96
97 definition option_cons ≝ λsig.λc:option sig.λl.
98   match c with [ None ⇒ l | Some c0 ⇒ c0::l ].
99
100 lemma opt_cons_tail_expand : ∀A,l.l = option_cons A (option_hd ? l) (tail ? l). 
101 #A * //
102 qed.
103
104 (* comparing lists *)
105
106 lemma compare_append : ∀A,l1,l2,l3,l4. l1@l2 = l3@l4 → 
107 ∃l:list A.(l1 = l3@l ∧ l4=l@l2) ∨ (l3 = l1@l ∧ l2=l@l4).
108 #A #l1 elim l1
109   [#l2 #l3 #l4 #Heq %{l3} %2 % // @Heq
110   |#a1 #tl1 #Hind #l2 #l3 cases l3
111     [#l4 #Heq %{(a1::tl1)} %1 % // @sym_eq @Heq 
112     |#a3 #tl3 #l4 normalize in ⊢ (%→?); #Heq cases (Hind l2 tl3 l4 ?)
113       [#l * * #Heq1 #Heq2 %{l}
114         [%1 % // >Heq1 >(cons_injective_l ????? Heq) //
115         |%2 % // >Heq1 >(cons_injective_l ????? Heq) //
116         ]
117       |@(cons_injective_r ????? Heq) 
118       ]
119     ]
120   ]
121 qed.
122 (**************************** iterators ******************************)
123
124 let rec map (A,B:Type[0]) (f: A → B) (l:list A) on l: list B ≝
125  match l with [ nil ⇒ nil ? | cons x tl ⇒ f x :: (map A B f tl)].
126
127 lemma map_append : ∀A,B,f,l1,l2.
128   (map A B f l1) @ (map A B f l2) = map A B f (l1@l2).
129 #A #B #f #l1 elim l1
130 [ #l2 @refl
131 | #h #t #IH #l2 normalize //
132 ] qed.
133   
134 let rec foldr (A,B:Type[0]) (f:A → B → B) (b:B) (l:list A) on l :B ≝  
135  match l with [ nil ⇒ b | cons a l ⇒ f a (foldr A B f b l)].
136  
137 definition filter ≝ 
138   λT.λp:T → bool.
139   foldr T (list T) (λx,l0.if p x then x::l0 else l0) (nil T).
140
141 (* compose f [a1;...;an] [b1;...;bm] = 
142   [f a1 b1; ... ;f an b1; ... ;f a1 bm; f an bm] *)
143  
144 definition compose ≝ λA,B,C.λf:A→B→C.λl1,l2.
145     foldr ?? (λi,acc.(map ?? (f i) l2)@acc) [ ] l1.
146
147 lemma filter_true : ∀A,l,a,p. p a = true → 
148   filter A p (a::l) = a :: filter A p l.
149 #A #l #a #p #pa (elim l) normalize >pa normalize // qed.
150
151 lemma filter_false : ∀A,l,a,p. p a = false → 
152   filter A p (a::l) = filter A p l.
153 #A #l #a #p #pa (elim l) normalize >pa normalize // qed.
154
155 theorem eq_map : ∀A,B,f,g,l. (∀x.f x = g x) → map A B f l = map A B g l.
156 #A #B #f #g #l #eqfg (elim l) normalize // qed.
157
158 (**************************** reverse *****************************)
159 let rec rev_append S (l1,l2:list S) on l1 ≝
160   match l1 with 
161   [ nil ⇒ l2
162   | cons a tl ⇒ rev_append S tl (a::l2)
163   ]
164 .
165
166 definition reverse ≝λS.λl.rev_append S l [].
167
168 lemma reverse_single : ∀S,a. reverse S [a] = [a]. 
169 // qed.
170
171 lemma rev_append_def : ∀S,l1,l2. 
172   rev_append S l1 l2 = (reverse S l1) @ l2 .
173 #S #l1 elim l1 normalize // 
174 qed.
175
176 lemma reverse_cons : ∀S,a,l. reverse S (a::l) = (reverse S l)@[a].
177 #S #a #l whd in ⊢ (??%?); // 
178 qed.
179
180 lemma reverse_append: ∀S,l1,l2. 
181   reverse S (l1 @ l2) = (reverse S l2)@(reverse S l1).
182 #S #l1 elim l1 [normalize // | #a #tl #Hind #l2 >reverse_cons
183 >reverse_cons // qed.
184
185 lemma reverse_reverse : ∀S,l. reverse S (reverse S l) = l.
186 #S #l elim l // #a #tl #Hind >reverse_cons >reverse_append 
187 normalize // qed.
188
189 (* an elimination principle for lists working on the tail;
190 useful for strings *)
191 lemma list_elim_left: ∀S.∀P:list S → Prop. P (nil S) →
192 (∀a.∀tl.P tl → P (tl@[a])) → ∀l. P l.
193 #S #P #Pnil #Pstep #l <(reverse_reverse … l) 
194 generalize in match (reverse S l); #l elim l //
195 #a #tl #H >reverse_cons @Pstep //
196 qed.
197
198 (**************************** length ******************************)
199
200 let rec length (A:Type[0]) (l:list A) on l ≝ 
201   match l with 
202     [ nil ⇒ 0
203     | cons a tl ⇒ S (length A tl)].
204
205 interpretation "list length" 'card l = (length ? l).
206
207 lemma length_tail: ∀A,l. length ? (tail A l) = pred (length ? l).
208 #A #l elim l // 
209 qed.
210
211 lemma length_tail1 : ∀A,l.0 < |l| → |tail A l| < |l|.
212 #A * normalize //
213 qed.
214
215 lemma length_append: ∀A.∀l1,l2:list A. 
216   |l1@l2| = |l1|+|l2|.
217 #A #l1 elim l1 // normalize /2/
218 qed.
219
220 lemma length_map: ∀A,B,l.∀f:A→B. length ? (map ?? f l) = length ? l.
221 #A #B #l #f elim l // #a #tl #Hind normalize //
222 qed.
223
224 lemma length_reverse: ∀A.∀l:list A. 
225   |reverse A l| = |l|.
226 #A #l elim l // #a #l0 #IH >reverse_cons >length_append normalize //
227 qed.
228
229 lemma lenght_to_nil: ∀A.∀l:list A.
230   |l| = 0 → l = [ ].
231 #A * // #a #tl normalize #H destruct
232 qed.
233  
234 lemma lists_length_split : 
235  ∀A.∀l1,l2:list A.(∃la,lb.(|la| = |l1| ∧ l2 = la@lb) ∨ (|la| = |l2| ∧ l1 = la@lb)).
236 #A #l1 elim l1
237 [ #l2 %{[ ]} %{l2} % % %
238 | #hd1 #tl1 #IH *
239   [ %{[ ]} %{(hd1::tl1)} %2 % %
240   | #hd2 #tl2 cases (IH tl2) #x * #y *
241     [ * #IH1 #IH2 %{(hd2::x)} %{y} % normalize % //
242     | * #IH1 #IH2 %{(hd1::x)} %{y} %2 normalize % // ]
243   ]
244 ]
245 qed.
246
247 (****************** traversing two lists in parallel *****************)
248 lemma list_ind2 : 
249   ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:list T1 → list T2 → Prop.
250   length ? l1 = length ? l2 →
251   (P [] []) → 
252   (∀tl1,tl2,hd1,hd2. P tl1 tl2 → P (hd1::tl1) (hd2::tl2)) → 
253   P l1 l2.
254 #T1 #T2 #l1 #l2 #P #Hl #Pnil #Pcons
255 generalize in match Hl; generalize in match l2;
256 elim l1
257 [#l2 cases l2 // normalize #t2 #tl2 #H destruct
258 |#t1 #tl1 #IH #l2 cases l2
259    [normalize #H destruct
260    |#t2 #tl2 #H @Pcons @IH normalize in H; destruct // ]
261 ]
262 qed.
263
264 lemma list_cases2 : 
265   ∀T1,T2:Type[0].∀l1:list T1.∀l2:list T2.∀P:Prop.
266   length ? l1 = length ? l2 →
267   (l1 = [] → l2 = [] → P) → 
268   (∀hd1,hd2,tl1,tl2.l1 = hd1::tl1 → l2 = hd2::tl2 → P) → P.
269 #T1 #T2 #l1 #l2 #P #Hl @(list_ind2 … Hl)
270 [ #Pnil #Pcons @Pnil //
271 | #tl1 #tl2 #hd1 #hd2 #IH1 #IH2 #Hp @Hp // ]
272 qed.
273
274 (*********************** properties of append ***********************)
275 lemma append_l1_injective : 
276   ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l1 = l2.
277 #a #l1 #l2 #l3 #l4 #Hlen @(list_ind2 … Hlen) //
278 #tl1 #tl2 #hd1 #hd2 #IH normalize #Heq destruct @eq_f /2/
279 qed.
280   
281 lemma append_l2_injective : 
282   ∀A.∀l1,l2,l3,l4:list A. |l1| = |l2| → l1@l3 = l2@l4 → l3 = l4.
283 #a #l1 #l2 #l3 #l4 #Hlen @(list_ind2 … Hlen) normalize //
284 #tl1 #tl2 #hd1 #hd2 #IH normalize #Heq destruct /2/
285 qed.
286
287 lemma append_l1_injective_r :
288   ∀A.∀l1,l2,l3,l4:list A. |l3| = |l4| → l1@l3 = l2@l4 → l1 = l2.
289 #a #l1 #l2 #l3 #l4 #Hlen #Heq lapply (eq_f … (reverse ?) … Heq)
290 >reverse_append >reverse_append #Heq1
291 lapply (append_l2_injective … Heq1) [ // ] #Heq2
292 lapply (eq_f … (reverse ?) … Heq2) //
293 qed.
294   
295 lemma append_l2_injective_r : 
296   ∀A.∀l1,l2,l3,l4:list A. |l3| = |l4| → l1@l3 = l2@l4 → l3 = l4.
297 #a #l1 #l2 #l3 #l4 #Hlen #Heq lapply (eq_f … (reverse ?) … Heq)
298 >reverse_append >reverse_append #Heq1
299 lapply (append_l1_injective … Heq1) [ // ] #Heq2
300 lapply (eq_f … (reverse ?) … Heq2) //
301 qed.
302
303 lemma length_rev_append: ∀A.∀l,acc:list A. 
304   |rev_append ? l acc| = |l|+|acc|.
305 #A #l elim l // #a #tl #Hind normalize 
306 #acc >Hind normalize // 
307 qed.
308
309 (****************************** mem ********************************)
310 let rec mem A (a:A) (l:list A) on l ≝
311   match l with
312   [ nil ⇒ False
313   | cons hd tl ⇒ a=hd ∨ mem A a tl
314   ]. 
315   
316 lemma mem_append: ∀A,a,l1,l2.mem A a (l1@l2) →
317   mem ? a l1 ∨ mem ? a l2.
318 #A #a #l1 elim l1 
319   [#l2 #mema %2 @mema
320   |#b #tl #Hind #l2 * 
321     [#eqab %1 %1 @eqab 
322     |#Hmema cases (Hind ? Hmema) -Hmema #Hmema [%1 %2 //|%2 //]
323     ]
324   ]
325 qed.
326
327 lemma mem_append_l1: ∀A,a,l1,l2.mem A a l1 → mem A a (l1@l2).
328 #A #a #l1 #l2 elim l1
329   [whd in ⊢ (%→?); @False_ind
330   |#b #tl #Hind * [#eqab %1 @eqab |#Hmema %2 @Hind //]
331   ]
332 qed.
333
334 lemma mem_append_l2: ∀A,a,l1,l2.mem A a l2 → mem A a (l1@l2).
335 #A #a #l1 #l2 elim l1 [//|#b #tl #Hind #Hmema %2 @Hind //]
336 qed.
337
338 lemma mem_single: ∀A,a,b. mem A a [b] → a=b.
339 #A #a #b * // @False_ind
340 qed.
341
342 lemma mem_map: ∀A,B.∀f:A→B.∀l,b. 
343   mem ? b (map … f l) → ∃a. mem ? a l ∧ f a = b.
344 #A #B #f #l elim l 
345   [#b normalize @False_ind
346   |#a #tl #Hind #b normalize *
347     [#eqb @(ex_intro … a) /3/
348     |#memb cases (Hind … memb) #a * #mema #eqb
349      @(ex_intro … a) /3/
350     ]
351   ]
352 qed.
353
354 lemma mem_map_forward: ∀A,B.∀f:A→B.∀a,l. 
355   mem A a l → mem B (f a) (map ?? f l).
356  #A #B #f #a #l elim l
357   [normalize @False_ind
358   |#b #tl #Hind * 
359     [#eqab <eqab normalize %1 % |#memtl normalize %2 @Hind @memtl]
360   ]
361 qed.
362
363 (***************************** split *******************************)
364 let rec split_rev A (l:list A) acc n on n ≝ 
365   match n with 
366   [O ⇒ 〈acc,l〉
367   |S m ⇒ match l with 
368     [nil ⇒ 〈acc,[]〉
369     |cons a tl ⇒ split_rev A tl (a::acc) m
370     ]
371   ].
372   
373 definition split ≝ λA,l,n.
374   let 〈l1,l2〉 ≝ split_rev A l [] n in 〈reverse ? l1,l2〉.
375
376 lemma split_rev_len: ∀A,n,l,acc. n ≤ |l| →
377   |\fst (split_rev A l acc n)| = n+|acc|.
378 #A #n elim n // #m #Hind *
379   [normalize #acc #Hfalse @False_ind /2/
380   |#a #tl #acc #Hlen normalize >Hind 
381     [normalize // |@le_S_S_to_le //]
382   ]
383 qed.
384
385 lemma split_len: ∀A,n,l. n ≤ |l| →
386   |\fst (split A l n)| = n.
387 #A #n #l #Hlen normalize >(eq_pair_fst_snd ?? (split_rev …))
388 normalize >length_reverse  >(split_rev_len … [ ] Hlen) normalize //
389 qed.
390   
391 lemma split_rev_eq: ∀A,n,l,acc. n ≤ |l| → 
392   reverse ? acc@ l = 
393     reverse ? (\fst (split_rev A l acc n))@(\snd (split_rev A l acc n)).
394  #A #n elim n //
395  #m #Hind * 
396    [#acc whd in ⊢ ((??%)→?); #False_ind /2/ 
397    |#a #tl #acc #Hlen >append_cons <reverse_single <reverse_append 
398     @(Hind tl) @le_S_S_to_le @Hlen
399    ]
400 qed.
401  
402 lemma split_eq: ∀A,n,l. n ≤ |l| → 
403   l = (\fst (split A l n))@(\snd (split A l n)).
404 #A #n #l #Hlen change with ((reverse ? [ ])@l) in ⊢ (??%?);
405 >(split_rev_eq … Hlen) normalize 
406 >(eq_pair_fst_snd ?? (split_rev A l [] n)) %
407 qed.
408
409 lemma split_exists: ∀A,n.∀l:list A. n ≤ |l| → 
410   ∃l1,l2. l = l1@l2 ∧ |l1| = n.
411 #A #n #l #Hlen @(ex_intro … (\fst (split A l n)))
412 @(ex_intro … (\snd (split A l n))) % /2/
413 qed.
414   
415 (****************************** flatten ******************************)
416 definition flatten ≝ λA.foldr (list A) (list A) (append A) [].
417
418 lemma flatten_to_mem: ∀A,n,l,l1,l2.∀a:list A. 0 < n →
419   (∀x. mem ? x l → |x| = n) → |a| = n → flatten ? l = l1@a@l2  →
420     (∃q.|l1| = n*q)  → mem ? a l.
421 #A #n #l elim l
422   [normalize #l1 #l2 #a #posn #Hlen #Ha #Hnil @False_ind
423    cut (|a|=0) [@sym_eq @le_n_O_to_eq 
424    @(transitive_le ? (|nil A|)) // >Hnil >length_append >length_append //] /2/
425   |#hd #tl #Hind #l1 #l2 #a #posn #Hlen #Ha 
426    whd in match (flatten ??); #Hflat * #q cases q
427     [<times_n_O #Hl1 
428      cut (a = hd) [>(lenght_to_nil… Hl1) in Hflat; 
429      whd in ⊢ ((???%)→?); #Hflat @sym_eq @(append_l1_injective … Hflat)
430      >Ha >Hlen // %1 //   
431      ] /2/
432     |#q1 #Hl1 lapply (split_exists … n l1 ?) //
433      * #l11 * #l12 * #Heql1 #Hlenl11 %2
434      @(Hind l12 l2 … posn ? Ha) 
435       [#x #memx @Hlen %2 //
436       |@(append_l2_injective ? hd l11) 
437         [>Hlenl11 @Hlen %1 %
438         |>Hflat >Heql1 >associative_append %
439         ]
440       |@(ex_intro …q1) @(injective_plus_r n) 
441        <Hlenl11 in ⊢ (??%?); <length_append <Heql1 >Hl1 //
442       ]
443     ]
444   ]
445 qed.
446
447 (****************************** nth ********************************)
448 let rec nth n (A:Type[0]) (l:list A) (d:A)  ≝  
449   match n with
450     [O ⇒ hd A l d
451     |S m ⇒ nth m A (tail A l) d].
452
453 lemma nth_nil: ∀A,a,i. nth i A ([]) a = a.
454 #A #a #i elim i normalize //
455 qed.
456
457 (****************************** nth_opt ********************************)
458 let rec nth_opt (A:Type[0]) (n:nat) (l:list A) on l : option A ≝
459 match l with
460 [ nil ⇒ None ?
461 | cons h t ⇒ match n with [ O ⇒ Some ? h | S m ⇒ nth_opt A m t ]
462 ].
463
464 (**************************** All *******************************)
465
466 let rec All (A:Type[0]) (P:A → Prop) (l:list A) on l : Prop ≝
467 match l with
468 [ nil ⇒ True
469 | cons h t ⇒ P h ∧ All A P t
470 ].
471
472 lemma All_mp : ∀A,P,Q. (∀a.P a → Q a) → ∀l. All A P l → All A Q l.
473 #A #P #Q #H #l elim l normalize //
474 #h #t #IH * /3/
475 qed.
476
477 lemma All_nth : ∀A,P,n,l.
478   All A P l →
479   ∀a.
480   nth_opt A n l = Some A a →
481   P a.
482 #A #P #n elim n
483 [ * [ #_ #a #E whd in E:(??%?); destruct
484     | #hd #tl * #H #_ #a #E whd in E:(??%?); destruct @H
485     ]
486 | #m #IH *
487   [ #_ #a #E whd in E:(??%?); destruct
488   | #hd #tl * #_ whd in ⊢ (? → ∀_.??%? → ?); @IH
489   ]
490 ] qed.
491
492 lemma All_append: ∀A,P,l1,l2. All A P l1 → All A P l2 → All A P (l1@l2).
493 #A #P #l1 elim l1 -l1 //
494 #a #l1 #IHl1 #l2 * /3 width=1/
495 qed.
496
497 lemma All_inv_append: ∀A,P,l1,l2. All A P (l1@l2) → All A P l1 ∧ All A P l2.
498 #A #P #l1 elim l1 -l1 /2 width=1/
499 #a #l1 #IHl1 #l2 * #Ha #Hl12
500 elim (IHl1 … Hl12) -IHl1 -Hl12 /3 width=1/
501 qed-.
502
503 (**************************** Allr ******************************)
504
505 let rec Allr (A:Type[0]) (R:relation A) (l:list A) on l : Prop ≝
506 match l with
507 [ nil       ⇒ True
508 | cons a1 l ⇒ match l with [ nil ⇒ True | cons a2 _ ⇒ R a1 a2 ∧ Allr A R l ]
509 ].
510
511 lemma Allr_fwd_append_sn: ∀A,R,l1,l2. Allr A R (l1@l2) → Allr A R l1.
512 #A #R #l1 elim l1 -l1 // #a1 * // #a2 #l1 #IHl1 #l2 * /3 width=2/
513 qed-.
514
515 lemma Allr_fwd_cons: ∀A,R,a,l. Allr A R (a::l) → Allr A R l.
516 #A #R #a * // #a0 #l * //
517 qed-.
518
519 lemma Allr_fwd_append_dx: ∀A,R,l1,l2. Allr A R (l1@l2) → Allr A R l2.
520 #A #R #l1 elim l1 -l1 // #a1 #l1 #IHl1 #l2 #H
521 lapply (Allr_fwd_cons … (l1@l2) H) -H /2 width=1/
522 qed-.  
523
524 (**************************** Exists *******************************)
525
526 let rec Exists (A:Type[0]) (P:A → Prop) (l:list A) on l : Prop ≝
527 match l with
528 [ nil ⇒ False
529 | cons h t ⇒ (P h) ∨ (Exists A P t)
530 ].
531
532 lemma Exists_append : ∀A,P,l1,l2.
533   Exists A P (l1 @ l2) → Exists A P l1 ∨ Exists A P l2.
534 #A #P #l1 elim l1
535 [ normalize /2/
536 | #h #t #IH #l2 *
537   [ #H /3/
538   | #H cases (IH l2 H) /3/
539   ]
540 ] qed.
541
542 lemma Exists_append_l : ∀A,P,l1,l2.
543   Exists A P l1 → Exists A P (l1@l2).
544 #A #P #l1 #l2 elim l1
545 [ *
546 | #h #t #IH *
547   [ #H %1 @H
548   | #H %2 @IH @H
549   ]
550 ] qed.
551
552 lemma Exists_append_r : ∀A,P,l1,l2.
553   Exists A P l2 → Exists A P (l1@l2).
554 #A #P #l1 #l2 elim l1
555 [ #H @H
556 | #h #t #IH #H %2 @IH @H
557 ] qed.
558
559 lemma Exists_add : ∀A,P,l1,x,l2. Exists A P (l1@l2) → Exists A P (l1@x::l2).
560 #A #P #l1 #x #l2 elim l1
561 [ normalize #H %2 @H
562 | #h #t #IH normalize * [ #H %1 @H | #H %2 @IH @H ]
563 qed.
564
565 lemma Exists_mid : ∀A,P,l1,x,l2. P x → Exists A P (l1@x::l2).
566 #A #P #l1 #x #l2 #H elim l1
567 [ %1 @H
568 | #h #t #IH %2 @IH
569 ] qed.
570
571 lemma Exists_map : ∀A,B,P,Q,f,l.
572 Exists A P l →
573 (∀a.P a → Q (f a)) →
574 Exists B Q (map A B f l).
575 #A #B #P #Q #f #l elim l //
576 #h #t #IH * [ #H #F %1 @F @H | #H #F %2 @IH [ @H | @F ] ] qed.
577
578 lemma Exists_All : ∀A,P,Q,l.
579   Exists A P l →
580   All A Q l →
581   ∃x. P x ∧ Q x.
582 #A #P #Q #l elim l [ * | #hd #tl #IH * [ #H1 * #H2 #_ %{hd} /2/ | #H1 * #_ #H2 @IH // ]
583 qed.
584
585 (**************************** fold *******************************)
586
587 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 ≝  
588  match l with 
589   [ nil ⇒ b 
590   | cons a l ⇒
591      if p a then op (f a) (fold A B op b p f l)
592      else fold A B op b p f l].
593       
594 notation "\fold  [ op , nil ]_{ ident i ∈ l | p} f"
595   with precedence 80
596 for @{'fold $op $nil (λ${ident i}. $p) (λ${ident i}. $f) $l}.
597
598 notation "\fold [ op , nil ]_{ident i ∈ l } f"
599   with precedence 80
600 for @{'fold $op $nil (λ${ident i}.true) (λ${ident i}. $f) $l}.
601
602 interpretation "\fold" 'fold op nil p f l = (fold ? ? op nil p f l).
603
604 theorem fold_true: 
605 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A→B. p a = true → 
606   \fold[op,nil]_{i ∈ a::l| p i} (f i) = 
607     op (f a) \fold[op,nil]_{i ∈ l| p i} (f i). 
608 #A #B #a #l #p #op #nil #f #pa normalize >pa // qed.
609
610 theorem fold_false: 
611 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f.
612 p a = false → \fold[op,nil]_{i ∈ a::l| p i} (f i) = 
613   \fold[op,nil]_{i ∈ l| p i} (f i).
614 #A #B #a #l #p #op #nil #f #pa normalize >pa // qed.
615
616 theorem fold_filter: 
617 ∀A,B.∀a:A.∀l.∀p.∀op:B→B→B.∀nil.∀f:A →B.
618   \fold[op,nil]_{i ∈ l| p i} (f i) = 
619     \fold[op,nil]_{i ∈ (filter A p l)} (f i).
620 #A #B #a #l #p #op #nil #f elim l //  
621 #a #tl #Hind cases(true_or_false (p a)) #pa 
622   [ >filter_true // > fold_true // >fold_true //
623   | >filter_false // >fold_false // ]
624 qed.
625
626 record Aop (A:Type[0]) (nil:A) : Type[0] ≝
627   {op :2> A → A → A; 
628    nill:∀a. op nil a = a; 
629    nilr:∀a. op a nil = a;
630    assoc: ∀a,b,c.op a (op b c) = op (op a b) c
631   }.
632
633 theorem fold_sum: ∀A,B. ∀I,J:list A.∀nil.∀op:Aop B nil.∀f.
634   op (\fold[op,nil]_{i∈I} (f i)) (\fold[op,nil]_{i∈J} (f i)) =
635     \fold[op,nil]_{i∈(I@J)} (f i).
636 #A #B #I #J #nil #op #f (elim I) normalize 
637   [>nill //|#a #tl #Hind <assoc //]
638 qed.
639
640 (********************** lhd and ltl ******************************)
641
642 let rec lhd (A:Type[0]) (l:list A) n on n ≝ match n with
643    [ O   ⇒ nil …
644    | S n ⇒ match l with [ nil ⇒ nil … | cons a l ⇒ a :: lhd A l n ]
645    ].
646
647 let rec ltl (A:Type[0]) (l:list A) n on n ≝ match n with
648    [ O   ⇒ l
649    | S n ⇒ ltl A (tail … l) n
650    ].
651
652 lemma lhd_nil: ∀A,n. lhd A ([]) n = [].
653 #A #n elim n //
654 qed.
655
656 lemma ltl_nil: ∀A,n. ltl A ([]) n = [].
657 #A #n elim n normalize //
658 qed.
659
660 lemma lhd_cons_ltl: ∀A,n,l. lhd A l n @ ltl A l n = l.
661 #A #n elim n -n //
662 #n #IHn #l elim l normalize //
663 qed.
664
665 lemma length_ltl: ∀A,n,l. |ltl A l n| = |l| - n.
666 #A #n elim n -n // 
667 #n #IHn *; normalize /2/
668 qed.
669
670 (********************** find ******************************)
671 let rec find (A,B:Type[0]) (f:A → option B) (l:list A) on l : option B ≝
672 match l with
673 [ nil ⇒ None B
674 | cons h t ⇒
675     match f h with
676     [ None ⇒ find A B f t
677     | Some b ⇒ Some B b
678     ]
679 ].
680
681 (********************** position_of ******************************)
682 let rec position_of_aux (A:Type[0]) (found: A → bool) (l:list A) (acc:nat) on l : option nat ≝
683 match l with
684 [ nil ⇒ None ?
685 | cons h t ⇒
686    match found h with [true ⇒ Some … acc | false ⇒ position_of_aux … found t (S acc)]].
687
688 definition position_of: ∀A:Type[0]. (A → bool) → list A → option nat ≝
689  λA,found,l. position_of_aux A found l 0.
690
691
692 (********************** make_list ******************************)
693 let rec make_list (A:Type[0]) (a:A) (n:nat) on n : list A ≝
694 match n with
695 [ O ⇒ [ ]
696 | S m ⇒ a::(make_list A a m)
697 ].