4 The fact of being able to decide, via a computable boolean function, the
5 equality between elements of a given set is an essential prerequisite for
6 effectively searching an element of that set inside a data structure. In this
7 section we shall define several boolean functions acting on lists of elements in
8 a DeqSet, and prove some of their properties.*)
10 include "basics/lists/list.ma".
11 include "tutorial/chapter4.ma".
13 (* The first function we define is an effective version of the membership relation,
14 between an element x and a list l. Its definition is a straightforward recursion on
17 let rec memb (S:DeqSet) (x:S) (l: list S) on l ≝
20 | cons a tl ⇒ (x == a) ∨ memb S x tl
23 notation < "\memb x l" non associative with precedence 90 for @{'memb $x $l}.
24 interpretation "boolean membership" 'memb a l = (memb ? a l).
26 (* We can now prove several interesing properties for memb:
27 - memb_hd: x is a member of x::l
28 - memb_cons: if x is a member of l than x is a member of a::l
29 - memb_single: if x is a member of [a] then x=a
30 - memb_append: if x is a member of l1@l2 then either x is a member of l1
31 or x is a member of l2
32 - memb_append_l1: if x is a member of l1 then x is a member of l1@l2
33 - memb_append_l2: if x is a member of l2 then x is a member of l1@l2
34 - memb_exists: if x is a member of l, than l can decomposed as l1@(x::l2)
35 - not_memb_to_not_eq: if x is not a member of l and y is, then x≠y
36 - memb_map: if a is a member of l, then (f a) is a member of (map f l)
37 - memb_compose: if a is a member of l1 and b is a meber of l2 than
38 (op a b) is a member of (compose op l1 l2)
41 lemma memb_hd: ∀S,a,l. memb S a (a::l) = true.
42 #S #a #l normalize >(proj2 … (eqb_true S …) (refl S a)) //
45 lemma memb_cons: ∀S,a,b,l.
46 memb S a l = true → memb S a (b::l) = true.
47 #S #a #b #l normalize cases (a==b) normalize //
50 lemma memb_single: ∀S,a,x. memb S a (x::[]) = true → a = x.
51 #S #a #x normalize cases (true_or_false … (a==x)) #H
52 [#_ >(\P H) // |>H normalize #abs @False_ind /2/]
55 lemma memb_append: ∀S,a,l1,l2.
56 memb S a (l1@l2) = true → memb S a l1= true ∨ memb S a l2 = true.
57 #S #a #l1 elim l1 normalize [#l2 #H %2 //]
58 #b #tl #Hind #l2 cases (a==b) normalize /2/
61 lemma memb_append_l1: ∀S,a,l1,l2.
62 memb S a l1= true → memb S a (l1@l2) = true.
63 #S #a #l1 elim l1 normalize
64 [normalize #le #abs @False_ind /2/
65 |#b #tl #Hind #l2 cases (a==b) normalize /2/
69 lemma memb_append_l2: ∀S,a,l1,l2.
70 memb S a l2= true → memb S a (l1@l2) = true.
71 #S #a #l1 elim l1 normalize //
72 #b #tl #Hind #l2 cases (a==b) normalize /2/
75 lemma memb_exists: ∀S,a,l.memb S a l = true → ∃l1,l2.l=l1@(a::l2).
76 #S #a #l elim l [normalize #abs @False_ind /2/]
77 #b #tl #Hind #H cases (orb_true_l … H)
78 [#eqba @(ex_intro … (nil S)) @(ex_intro … tl) >(\P eqba) //
79 |#mem_tl cases (Hind mem_tl) #l1 * #l2 #eqtl
80 @(ex_intro … (b::l1)) @(ex_intro … l2) >eqtl //
84 lemma not_memb_to_not_eq: ∀S,a,b,l.
85 memb S a l = false → memb S b l = true → a==b = false.
86 #S #a #b #l cases (true_or_false (a==b)) //
87 #eqab >(\P eqab) #H >H #abs @False_ind /2/
90 lemma memb_map: ∀S1,S2,f,a,l. memb S1 a l= true →
91 memb S2 (f a) (map … f l) = true.
92 #S1 #S2 #f #a #l elim l normalize [//]
93 #x #tl #memba cases (true_or_false (a==x))
94 [#eqx >eqx >(\P eqx) >(\b (refl … (f x))) normalize //
95 |#eqx >eqx cases (f a==f x) normalize /2/
99 lemma memb_compose: ∀S1,S2,S3,op,a1,a2,l1,l2.
100 memb S1 a1 l1 = true → memb S2 a2 l2 = true →
101 memb S3 (op a1 a2) (compose S1 S2 S3 op l1 l2) = true.
102 #S1 #S2 #S3 #op #a1 #a2 #l1 elim l1 [normalize //]
103 #x #tl #Hind #l2 #memba1 #memba2 cases (orb_true_l … memba1)
104 [#eqa1 >(\P eqa1) @memb_append_l1 @memb_map //
105 |#membtl @memb_append_l2 @Hind //
112 If we are interested in representing finite sets as lists, is is convenient
113 to avoid duplications of elements. The following uniqueb check this property.
116 let rec uniqueb (S:DeqSet) l on l : bool ≝
119 | cons a tl ⇒ notb (memb S a tl) ∧ uniqueb S tl
122 (* unique_append l1 l2 add l1 in fornt of l2, but preserving unicity *)
124 let rec unique_append (S:DeqSet) (l1,l2: list S) on l1 ≝
128 let r ≝ unique_append S tl l2 in
129 if memb S a r then r else a::r
132 lemma memb_unique_append: ∀S,a,l1,l2.
133 memb S a (unique_append S l1 l2) = true → memb S a l1= true ∨ memb S a l2 = true.
134 #S #a #l1 elim l1 normalize [#l2 #H %2 //]
135 #b #tl #Hind #l2 cases (true_or_false … (a==b)) #Hab >Hab normalize /2/
136 cases (memb S b (unique_append S tl l2)) normalize
137 [@Hind | >Hab normalize @Hind]
140 lemma unique_append_elim: ∀S:DeqSet.∀P: S → Prop.∀l1,l2.
141 (∀x. memb S x l1 = true → P x) → (∀x. memb S x l2 = true → P x) →
142 ∀x. memb S x (unique_append S l1 l2) = true → P x.
143 #S #P #l1 #l2 #Hl1 #Hl2 #x #membx cases (memb_unique_append … membx) /2/
146 lemma unique_append_unique: ∀S,l1,l2. uniqueb S l2 = true →
147 uniqueb S (unique_append S l1 l2) = true.
148 #S #l1 elim l1 normalize // #a #tl #Hind #l2 #uniquel2
149 cases (true_or_false … (memb S a (unique_append S tl l2)))
150 #H >H normalize [@Hind //] >H normalize @Hind //
158 λS,l1,l2.∀a. memb S a l1 = true → memb S a l2 = true.
160 lemma sublist_length: ∀S,l1,l2.
161 uniqueb S l1 = true → sublist S l1 l2 → |l1| ≤ |l2|.
163 #a #tl #Hind #l2 #unique #sub
164 cut (∃l3,l4.l2=l3@(a::l4)) [@memb_exists @sub //]
165 * #l3 * #l4 #eql2 >eql2 >length_append normalize
166 applyS le_S_S <length_append @Hind [@(andb_true_r … unique)]
167 >eql2 in sub; #sub #x #membx
168 cases (memb_append … (sub x (orb_true_r2 … membx)))
169 [#membxl3 @memb_append_l1 //
170 |#membxal4 cases (orb_true_l … membxal4)
171 [#eqxa @False_ind lapply (andb_true_l … unique)
172 <(\P eqxa) >membx normalize /2/ |#membxl4 @memb_append_l2 //
177 lemma sublist_unique_append_l1:
178 ∀S,l1,l2. sublist S l1 (unique_append S l1 l2).
179 #S #l1 elim l1 normalize [#l2 #S #abs @False_ind /2/]
181 normalize cases (true_or_false … (a==x)) #eqax >eqax
182 [<(\P eqax) cases (true_or_false (memb S a (unique_append S tl l2)))
183 [#H >H normalize // | #H >H normalize >(\b (refl … a)) //]
184 |cases (memb S x (unique_append S tl l2)) normalize
185 [/2/ |>eqax normalize /2/]
189 lemma sublist_unique_append_l2:
190 ∀S,l1,l2. sublist S l2 (unique_append S l1 l2).
191 #S #l1 elim l1 [normalize //] #x #tl #Hind normalize
192 #l2 #a cases (memb S x (unique_append S tl l2)) normalize
193 [@Hind | cases (a==x) normalize // @Hind]
196 lemma decidable_sublist:∀S,l1,l2.
197 (sublist S l1 l2) ∨ ¬(sublist S l1 l2).
199 [%1 #a normalize in ⊢ (%→?); #abs @False_ind /2/
201 [cases (true_or_false (memb S a l2)) #memba
202 [%1 whd #x #membx cases (orb_true_l … membx)
203 [#eqax >(\P eqax) // |@subtl]
204 |%2 @(not_to_not … (eqnot_to_noteq … true memba)) #H1 @H1 @memb_hd
206 |%2 @(not_to_not … subtl) #H1 #x #H2 @H1 @memb_cons //
215 lemma memb_filter_true: ∀S,f,a,l.
216 memb S a (filter S f l) = true → f a = true.
217 #S #f #a #l elim l [normalize #H @False_ind /2/]
218 #b #tl #Hind cases (true_or_false (f b)) #H
219 normalize >H normalize [2:@Hind]
220 cases (true_or_false (a==b)) #eqab
221 [#_ >(\P eqab) // | >eqab normalize @Hind]
224 lemma memb_filter_memb: ∀S,f,a,l.
225 memb S a (filter S f l) = true → memb S a l = true.
226 #S #f #a #l elim l [normalize //]
227 #b #tl #Hind normalize (cases (f b)) normalize
228 cases (a==b) normalize // @Hind
231 lemma memb_filter: ∀S,f,l,x. memb S x (filter ? f l) = true →
232 memb S x l = true ∧ (f x = true).
235 lemma memb_filter_l: ∀S,f,x,l. (f x = true) → memb S x l = true →
236 memb S x (filter ? f l) = true.
237 #S #f #x #l #fx elim l normalize //
238 #b #tl #Hind cases (true_or_false (x==b)) #eqxb
239 [<(\P eqxb) >(\b (refl … x)) >fx normalize >(\b (refl … x)) normalize //
240 |>eqxb cases (f b) normalize [>eqxb normalize @Hind| @Hind]
249 let rec exists (A:Type[0]) (p:A → bool) (l:list A) on l : bool ≝
252 | cons h t ⇒ orb (p h) (exists A p t)