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.
8 \ / This file is distributed under the terms of the
9 \ / GNU General Public License Version 2
10 V_______________________________________________________________ *)
13 (* boolean functions over lists *)
15 include "basics/lists/list.ma".
16 include "basics/sets.ma".
17 include "basics/deqsets.ma".
19 (********* isnilb *********)
20 let rec isnilb A (l: list A) on l ≝
26 (********* search *********)
28 let rec memb (S:DeqSet) (x:S) (l: list S) on l ≝
31 | cons a tl ⇒ (x == a) ∨ memb S x tl
34 interpretation "boolean membership" 'mem a l = (memb ? a l).
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)) //
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 //
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/]
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/
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/
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/
71 lemma memb_exists: ∀S,a,l.memb S a l = true →
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 //
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/
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/
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 //
106 lemma memb_reverse: ∀S:DeqSet.∀a:S.∀l.
107 memb ? a l = true → memb ? a (reverse ? l) = true.
108 #S #a #l elim l [normalize //]
109 #b #tl #Hind #memba change with ([b]@tl) in match (b::tl);
110 >reverse_append cases (orb_true_l … memba) #Hcase
111 [@memb_append_l2 >(\P Hcase) whd in match (reverse ??); @memb_hd
116 lemma mem_to_memb: ∀S:DeqSet.∀a,l. mem S a l → memb S a l = true.
117 #S #a #l elim l normalize
121 |#Hmem cases (a==hd) // normalize /2/
126 lemma memb_to_mem: ∀S:DeqSet.∀l,a. memb S a l =true → mem S a l.
128 [normalize #H destruct
129 |#b #tl #Hind #mema cases (orb_true_l … mema)
130 [#eqab >(\P eqab) %1 % |#memtl %2 @Hind @memtl]
134 (**************** unicity test *****************)
136 let rec uniqueb (S:DeqSet) l on l : bool ≝
139 | cons a tl ⇒ notb (memb S a tl) ∧ uniqueb S tl
142 (* unique_append l1 l2 add l1 in fornt of l2, but preserving unicity *)
144 let rec unique_append (S:DeqSet) (l1,l2: list S) on l1 ≝
148 let r ≝ unique_append S tl l2 in
149 if memb S a r then r else a::r
152 lemma memb_unique_append: ∀S,a,l1,l2.
153 memb S a (unique_append S l1 l2) = true →
154 memb S a l1= true ∨ memb S a l2 = true.
155 #S #a #l1 elim l1 normalize [#l2 #H %2 //]
156 #b #tl #Hind #l2 cases (true_or_false … (a==b)) #Hab >Hab normalize /2/
157 cases (memb S b (unique_append S tl l2)) normalize
158 [@Hind | >Hab normalize @Hind]
161 lemma unique_append_elim: ∀S:DeqSet.∀P: S → Prop.∀l1,l2.
162 (∀x. memb S x l1 = true → P x) → (∀x. memb S x l2 = true → P x) →
163 ∀x. memb S x (unique_append S l1 l2) = true → P x.
164 #S #P #l1 #l2 #Hl1 #Hl2 #x #membx cases (memb_unique_append … membx)
168 lemma unique_append_unique: ∀S,l1,l2. uniqueb S l2 = true →
169 uniqueb S (unique_append S l1 l2) = true.
170 #S #l1 elim l1 normalize // #a #tl #Hind #l2 #uniquel2
171 cases (true_or_false … (memb S a (unique_append S tl l2)))
172 #H >H normalize [@Hind //] >H normalize @Hind //
175 lemma memb_map_inj: ∀A,B:DeqSet.∀f:A→B.∀l,a. injective A B f →
176 memb ? (f a) (map ?? f l) = true → memb ? a l = true.
177 #A #B #f #l #a #injf elim l
179 |#a1 #tl #Hind #Ha cases (orb_true_l … Ha)
180 [#eqf @orb_true_r1 @(\b ?) >(injf … (\P eqf)) //
181 |#membtl @orb_true_r2 /2/
186 lemma unique_map_inj: ∀A,B:DeqSet.∀f:A→B.∀l. injective A B f →
187 uniqueb A l = true → uniqueb B (map ?? f l) = true .
188 #A #B #f #l #injf elim l
190 |#a #tl #Hind #Htl @true_to_andb_true
191 [@sym_eq @noteq_to_eqnot @sym_not_eq
192 @(not_to_not ?? (memb_map_inj … injf …) )
193 <(andb_true_l ?? Htl) /2/
194 |@Hind @(andb_true_r ?? Htl)
199 (******************* sublist *******************)
201 λS,l1,l2.∀a. memb S a l1 = true → memb S a l2 = true.
203 lemma sublist_length: ∀S,l1,l2.
204 uniqueb S l1 = true → sublist S l1 l2 → |l1| ≤ |l2|.
206 #a #tl #Hind #l2 #unique #sub
207 cut (∃l3,l4.l2=l3@(a::l4)) [@memb_exists @sub //]
208 * #l3 * #l4 #eql2 >eql2 >length_append normalize
209 applyS le_S_S <length_append @Hind [@(andb_true_r … unique)]
210 >eql2 in sub; #sub #x #membx
211 cases (memb_append … (sub x (orb_true_r2 … membx)))
212 [#membxl3 @memb_append_l1 //
213 |#membxal4 cases (orb_true_l … membxal4)
214 [#eqxa @False_ind lapply (andb_true_l … unique)
215 <(\P eqxa) >membx normalize /2/ |#membxl4 @memb_append_l2 //
220 lemma sublist_unique_append_l1:
221 ∀S,l1,l2. sublist S l1 (unique_append S l1 l2).
222 #S #l1 elim l1 normalize [#l2 #S #abs @False_ind /2/]
224 normalize cases (true_or_false … (a==x)) #eqax >eqax
225 [<(\P eqax) cases (true_or_false (memb S a (unique_append S tl l2)))
226 [#H >H normalize // | #H >H normalize >(\b (refl … a)) //]
227 |cases (memb S x (unique_append S tl l2)) normalize
228 [/2/ |>eqax normalize /2/]
232 lemma sublist_unique_append_l2:
233 ∀S,l1,l2. sublist S l2 (unique_append S l1 l2).
234 #S #l1 elim l1 [normalize //] #x #tl #Hind normalize
235 #l2 #a cases (memb S x (unique_append S tl l2)) normalize
236 [@Hind | cases (a==x) normalize // @Hind]
239 lemma decidable_sublist:∀S,l1,l2.
240 (sublist S l1 l2) ∨ ¬(sublist S l1 l2).
242 [%1 #a normalize in ⊢ (%→?); #abs @False_ind /2/
244 [cases (true_or_false (memb S a l2)) #memba
245 [%1 whd #x #membx cases (orb_true_l … membx)
246 [#eqax >(\P eqax) // |@subtl]
247 |%2 @(not_to_not … (eqnot_to_noteq … true memba)) #H1 @H1 @memb_hd
249 |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons //
254 (********************* filtering *****************)
256 lemma memb_filter_memb: ∀S,f,a,l.
257 memb S a (filter S f l) = true → memb S a l = true.
258 #S #f #a #l elim l [normalize //]
259 #b #tl #Hind normalize (cases (f b)) normalize
260 cases (a==b) normalize // @Hind
263 lemma uniqueb_filter : ∀S,l,f.
264 uniqueb S l = true → uniqueb S (filter S f l) = true.
266 #a #tl #Hind #Hunique cases (andb_true … Hunique)
267 #memba #uniquetl cases (true_or_false … (f a)) #Hfa
268 [>filter_true // @true_to_andb_true
269 [@sym_eq @noteq_to_eqnot @(not_to_not … (eqnot_to_noteq … (sym_eq … memba)))
270 #H @sym_eq @(memb_filter_memb … f) <H //
277 lemma filter_true: ∀S,f,a,l.
278 memb S a (filter S f l) = true → f a = true.
279 #S #f #a #l elim l [normalize #H @False_ind /2/]
280 #b #tl #Hind cases (true_or_false (f b)) #H
281 normalize >H normalize [2:@Hind]
282 cases (true_or_false (a==b)) #eqab
283 [#_ >(\P eqab) // | >eqab normalize @Hind]
286 lemma memb_filter: ∀S,f,l,x. memb S x (filter ? f l) = true →
287 memb S x l = true ∧ (f x = true).
290 lemma memb_filter_l: ∀S,f,x,l. (f x = true) → memb S x l = true →
291 memb S x (filter ? f l) = true.
292 #S #f #x #l #fx elim l normalize //
293 #b #tl #Hind cases (true_or_false (x==b)) #eqxb
294 [<(\P eqxb) >(\b (refl … x)) >fx normalize >(\b (refl … x)) normalize //
295 |>eqxb cases (f b) normalize [>eqxb normalize @Hind| @Hind]
299 (********************* exists *****************)
301 let rec exists (A:Type[0]) (p:A → bool) (l:list A) on l : bool ≝
304 | cons h t ⇒ orb (p h) (exists A p t)
307 lemma Exists_exists : ∀A,P,l.
310 #A #P #l elim l [ * | #hd #tl #IH * [ #H %{hd} @H | @IH ]