]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/basics/lists/listb.ma
FinOpt
[helm.git] / matita / matita / lib / basics / lists / listb.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
13 (* boolean functions over lists *)
14
15 include "basics/lists/list.ma".
16 include "basics/sets.ma".
17 include "basics/deqsets.ma".
18
19 (********* isnilb *********)
20 let rec isnilb A (l: list A) on l ≝
21 match l with
22 [ nil ⇒ true
23 | cons hd tl ⇒ false
24 ].
25
26 (********* search *********)
27
28 let rec memb (S:DeqSet) (x:S) (l: list S) on l  ≝
29   match l with
30   [ nil ⇒ false
31   | cons a tl ⇒ (x == a) ∨ memb S x tl
32   ].
33
34 interpretation "boolean membership" 'mem a l = (memb ? a l).
35
36 lemma memb_hd: ∀S,a,l. memb S a (a::l) = true.
37 #S #a #l normalize >(proj2 … (eqb_true S …) (refl S a)) //
38 qed.
39
40 lemma memb_cons: ∀S,a,b,l. 
41   memb S a l = true → memb S a (b::l) = true.
42 #S #a #b #l normalize cases (a==b) normalize // 
43 qed.
44
45 lemma memb_single: ∀S,a,x. memb S a [x] = true → a = x.
46 #S #a #x normalize cases (true_or_false … (a==x)) #H
47   [#_ >(\P H) // |>H normalize #abs @False_ind /2/]
48 qed.
49
50 lemma memb_append: ∀S,a,l1,l2. 
51 memb S a (l1@l2) = true →
52   memb S a l1= true ∨ memb S a l2 = true.
53 #S #a #l1 elim l1 normalize [#l2 #H %2 //] 
54 #b #tl #Hind #l2 cases (a==b) normalize /2/ 
55 qed. 
56
57 lemma memb_append_l1: ∀S,a,l1,l2. 
58  memb S a l1= true → memb S a (l1@l2) = true.
59 #S #a #l1 elim l1 normalize
60   [normalize #le #abs @False_ind /2/
61   |#b #tl #Hind #l2 cases (a==b) normalize /2/ 
62   ]
63 qed. 
64
65 lemma memb_append_l2: ∀S,a,l1,l2. 
66  memb S a l2= true → memb S a (l1@l2) = true.
67 #S #a #l1 elim l1 normalize //
68 #b #tl #Hind #l2 cases (a==b) normalize /2/ 
69 qed. 
70
71 lemma memb_exists: ∀S,a,l.memb S a l = true → 
72   ∃l1,l2.l=l1@(a::l2).
73 #S #a #l elim l [normalize #abs @False_ind /2/]
74 #b #tl #Hind #H cases (orb_true_l … H)
75   [#eqba @(ex_intro … (nil S)) @(ex_intro … tl) >(\P eqba) //
76   |#mem_tl cases (Hind mem_tl) #l1 * #l2 #eqtl
77    @(ex_intro … (b::l1)) @(ex_intro … l2) >eqtl //
78   ]
79 qed.
80
81 lemma not_memb_to_not_eq: ∀S,a,b,l. 
82  memb S a l = false → memb S b l = true → a==b = false.
83 #S #a #b #l cases (true_or_false (a==b)) // 
84 #eqab >(\P eqab) #H >H #abs @False_ind /2/
85 qed. 
86  
87 lemma memb_map: ∀S1,S2,f,a,l. memb S1 a l= true → 
88   memb S2 (f a) (map … f l) = true.
89 #S1 #S2 #f #a #l elim l normalize [//]
90 #x #tl #memba cases (true_or_false (a==x))
91   [#eqx >eqx >(\P eqx) >(\b (refl … (f x))) normalize //
92   |#eqx >eqx cases (f a==f x) normalize /2/
93   ]
94 qed.
95
96 lemma memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2.   
97   memb S1 a1 l1 = true → memb S2 a2 l2 = true →
98   memb S3 (op a1 a2) (compose S1 S2 S3 op l1 l2) = true.
99 #S1 #S2 #S3 #op #a1 #a2 #l1 elim l1 [normalize //]
100 #x #tl #Hind #l2 #memba1 #memba2 cases (orb_true_l … memba1)
101   [#eqa1 >(\P eqa1) @memb_append_l1 @memb_map // 
102   |#membtl @memb_append_l2 @Hind //
103   ]
104 qed.
105
106 lemma mem_to_memb: ∀S:DeqSet.∀a,l. mem S a l → memb S a l = true.
107 #S #a #l elim l normalize
108   [@False_ind
109   |#hd #tl #Hind *
110     [#eqa >(\b eqa) %
111     |#Hmem cases (a==hd) // normalize /2/
112     ]
113   ]
114 qed.
115 (**************** unicity test *****************)
116
117 let rec uniqueb (S:DeqSet) l on l : bool ≝
118   match l with 
119   [ nil ⇒ true
120   | cons a tl ⇒ notb (memb S a tl) ∧ uniqueb S tl
121   ].
122
123 (* unique_append l1 l2 add l1 in fornt of l2, but preserving unicity *)
124
125 let rec unique_append (S:DeqSet) (l1,l2: list S) on l1 ≝
126   match l1 with
127   [ nil ⇒ l2
128   | cons a tl ⇒ 
129      let r ≝ unique_append S tl l2 in
130      if memb S a r then r else a::r
131   ].
132
133 lemma memb_unique_append: ∀S,a,l1,l2. 
134 memb S a (unique_append S l1 l2) = true →
135   memb S a l1= true ∨ memb S a l2 = true.
136 #S #a #l1 elim l1 normalize [#l2 #H %2 //] 
137 #b #tl #Hind #l2 cases (true_or_false … (a==b)) #Hab >Hab normalize /2/
138 cases (memb S b (unique_append S tl l2)) normalize 
139   [@Hind | >Hab normalize @Hind]   
140 qed. 
141
142 lemma unique_append_elim: ∀S:DeqSet.∀P: S → Prop.∀l1,l2. 
143 (∀x. memb S x l1 = true → P x) → (∀x. memb S x l2 = true → P x) →
144 ∀x. memb S x (unique_append S l1 l2) = true → P x. 
145 #S #P #l1 #l2 #Hl1 #Hl2 #x #membx cases (memb_unique_append … membx)
146 /2/ 
147 qed.
148
149 lemma unique_append_unique: ∀S,l1,l2. uniqueb S l2 = true →
150   uniqueb S (unique_append S l1 l2) = true.
151 #S #l1 elim l1 normalize // #a #tl #Hind #l2 #uniquel2
152 cases (true_or_false … (memb S a (unique_append S tl l2))) 
153 #H >H normalize [@Hind //] >H normalize @Hind //
154 qed.
155
156 lemma memb_map_inj: ∀A,B:DeqSet.∀f:A→B.∀l,a. injective A B f → 
157   memb ? (f a) (map ?? f l) = true → memb ? a l = true.
158 #A #B #f #l #a #injf elim l
159   [normalize //
160   |#a1 #tl #Hind #Ha cases (orb_true_l … Ha)
161     [#eqf @orb_true_r1 @(\b ?)  >(injf … (\P eqf)) //
162     |#membtl @orb_true_r2 /2/
163     ]
164   ]
165 qed.
166
167 lemma unique_map_inj: ∀A,B:DeqSet.∀f:A→B.∀l. injective A B f → 
168   uniqueb A l = true → uniqueb B (map ?? f l) = true .
169 #A #B #f #l #injf elim l 
170   [normalize //
171   |#a #tl #Hind #Htl @true_to_andb_true
172     [@sym_eq @noteq_to_eqnot @sym_not_eq 
173      @(not_to_not ?? (memb_map_inj … injf …) )
174      <(andb_true_l ?? Htl) /2/
175     |@Hind @(andb_true_r ?? Htl)
176     ]
177   ]
178 qed.
179
180 (******************* sublist *******************)
181 definition sublist ≝ 
182   λS,l1,l2.∀a. memb S a l1 = true → memb S a l2 = true.
183
184 lemma sublist_length: ∀S,l1,l2. 
185  uniqueb S l1 = true → sublist S l1 l2 → |l1| ≤ |l2|.
186 #S #l1 elim l1 // 
187 #a #tl #Hind #l2 #unique #sub
188 cut (∃l3,l4.l2=l3@(a::l4)) [@memb_exists @sub //]
189 * #l3 * #l4 #eql2 >eql2 >length_append normalize 
190 applyS le_S_S <length_append @Hind [@(andb_true_r … unique)]
191 >eql2 in sub; #sub #x #membx 
192 cases (memb_append … (sub x (orb_true_r2 … membx)))
193   [#membxl3 @memb_append_l1 //
194   |#membxal4 cases (orb_true_l … membxal4)
195     [#eqxa @False_ind lapply (andb_true_l … unique)
196      <(\P eqxa) >membx normalize /2/ |#membxl4 @memb_append_l2 //
197     ]
198   ]
199 qed.
200
201 lemma sublist_unique_append_l1: 
202   ∀S,l1,l2. sublist S l1 (unique_append S l1 l2).
203 #S #l1 elim l1 normalize [#l2 #S #abs @False_ind /2/]
204 #x #tl #Hind #l2 #a 
205 normalize cases (true_or_false … (a==x)) #eqax >eqax 
206 [<(\P eqax) cases (true_or_false (memb S a (unique_append S tl l2)))
207   [#H >H normalize // | #H >H normalize >(\b (refl … a)) //]
208 |cases (memb S x (unique_append S tl l2)) normalize 
209   [/2/ |>eqax normalize /2/]
210 ]
211 qed.
212
213 lemma sublist_unique_append_l2: 
214   ∀S,l1,l2. sublist S l2 (unique_append S l1 l2).
215 #S #l1 elim l1 [normalize //] #x #tl #Hind normalize 
216 #l2 #a cases (memb S x (unique_append S tl l2)) normalize
217 [@Hind | cases (a==x) normalize // @Hind]
218 qed.
219
220 lemma decidable_sublist:∀S,l1,l2. 
221   (sublist S l1 l2) ∨ ¬(sublist S l1 l2).
222 #S #l1 #l2 elim l1 
223   [%1 #a normalize in ⊢ (%→?); #abs @False_ind /2/
224   |#a #tl * #subtl 
225     [cases (true_or_false (memb S a l2)) #memba
226       [%1 whd #x #membx cases (orb_true_l … membx)
227         [#eqax >(\P eqax) // |@subtl]
228       |%2 @(not_to_not … (eqnot_to_noteq … true memba)) #H1 @H1 @memb_hd
229       ]
230     |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons //
231     ] 
232   ]
233 qed.
234
235 (********************* filtering *****************)
236
237 lemma memb_filter_memb: ∀S,f,a,l. 
238   memb S a (filter S f l) = true → memb S a l = true.
239 #S #f #a #l elim l [normalize //]
240 #b #tl #Hind normalize (cases (f b)) normalize 
241 cases (a==b) normalize // @Hind
242 qed.
243   
244 lemma uniqueb_filter : ∀S,l,f.
245  uniqueb S l = true → uniqueb S (filter S f l) = true.
246 #S #l #f elim l //
247 #a #tl #Hind #Hunique cases (andb_true … Hunique)
248 #memba #uniquetl cases (true_or_false … (f a)) #Hfa
249   [>filter_true // @true_to_andb_true
250     [@sym_eq @noteq_to_eqnot @(not_to_not … (eqnot_to_noteq … (sym_eq … memba)))
251      #H @sym_eq @(memb_filter_memb … f) <H // 
252     |/2/
253     ]
254   |>filter_false /2/
255   ]
256 qed.
257   
258 lemma filter_true: ∀S,f,a,l. 
259   memb S a (filter S f l) = true → f a = true.
260 #S #f #a #l elim l [normalize #H @False_ind /2/]
261 #b #tl #Hind cases (true_or_false (f b)) #H
262 normalize >H normalize [2:@Hind]
263 cases (true_or_false (a==b)) #eqab
264   [#_ >(\P eqab) // | >eqab normalize @Hind]
265 qed. 
266   
267 lemma memb_filter: ∀S,f,l,x. memb S x (filter ? f l) = true → 
268 memb S x l = true ∧ (f x = true).
269 /3/ qed.
270
271 lemma memb_filter_l: ∀S,f,x,l. (f x = true) → memb S x l = true →
272 memb S x (filter ? f l) = true.
273 #S #f #x #l #fx elim l normalize //
274 #b #tl #Hind cases (true_or_false (x==b)) #eqxb
275   [<(\P eqxb) >(\b (refl … x)) >fx normalize >(\b (refl … x)) normalize //
276   |>eqxb cases (f b) normalize [>eqxb normalize @Hind| @Hind]
277   ]
278 qed. 
279
280 (********************* exists *****************)
281
282 let rec exists (A:Type[0]) (p:A → bool) (l:list A) on l : bool ≝
283 match l with
284 [ nil ⇒ false
285 | cons h t ⇒ orb (p h) (exists A p t)
286 ].
287
288 lemma Exists_exists : ∀A,P,l.
289   Exists A P l →
290   ∃x. P x.
291 #A #P #l elim l [ * | #hd #tl #IH * [ #H %{hd} @H | @IH ]
292 qed.