1 (* Title: HOL/Quotient_Examples/FSet.thy |
|
2 Author: Cezary Kaliszyk, TU Munich |
|
3 Author: Christian Urban, TU Munich |
|
4 |
|
5 Type of finite sets. |
|
6 *) |
|
7 |
|
8 theory FSet |
|
9 imports Quotient_List |
|
10 begin |
|
11 |
|
12 text {* |
|
13 The type of finite sets is created by a quotient construction |
|
14 over lists. The definition of the equivalence: |
|
15 *} |
|
16 |
|
17 fun |
|
18 list_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" (infix "\<approx>" 50) |
|
19 where |
|
20 "list_eq xs ys \<longleftrightarrow> set xs = set ys" |
|
21 |
|
22 lemma list_eq_equivp: |
|
23 shows "equivp list_eq" |
|
24 unfolding equivp_reflp_symp_transp |
|
25 unfolding reflp_def symp_def transp_def |
|
26 by auto |
|
27 |
|
28 text {* Fset type *} |
|
29 |
|
30 quotient_type |
|
31 'a fset = "'a list" / "list_eq" |
|
32 by (rule list_eq_equivp) |
|
33 |
|
34 text {* |
|
35 Definitions for membership, sublist, cardinality, |
|
36 intersection, difference and respectful fold over |
|
37 lists. |
|
38 *} |
|
39 |
|
40 definition |
|
41 memb :: "'a \<Rightarrow> 'a list \<Rightarrow> bool" |
|
42 where |
|
43 [simp]: "memb x xs \<longleftrightarrow> x \<in> set xs" |
|
44 |
|
45 definition |
|
46 sub_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" |
|
47 where |
|
48 [simp]: "sub_list xs ys \<longleftrightarrow> set xs \<subseteq> set ys" |
|
49 |
|
50 definition |
|
51 card_list :: "'a list \<Rightarrow> nat" |
|
52 where |
|
53 [simp]: "card_list xs = card (set xs)" |
|
54 |
|
55 definition |
|
56 inter_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
57 where |
|
58 [simp]: "inter_list xs ys = [x \<leftarrow> xs. x \<in> set xs \<and> x \<in> set ys]" |
|
59 |
|
60 definition |
|
61 diff_list :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
62 where |
|
63 [simp]: "diff_list xs ys = [x \<leftarrow> xs. x \<notin> set ys]" |
|
64 |
|
65 definition |
|
66 rsp_fold |
|
67 where |
|
68 "rsp_fold f \<equiv> \<forall>u v w. (f u (f v w) = f v (f u w))" |
|
69 |
|
70 primrec |
|
71 fold_list :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a list \<Rightarrow> 'b" |
|
72 where |
|
73 "fold_list f z [] = z" |
|
74 | "fold_list f z (a # xs) = |
|
75 (if (rsp_fold f) then |
|
76 if a \<in> set xs then fold_list f z xs |
|
77 else f a (fold_list f z xs) |
|
78 else z)" |
|
79 |
|
80 |
|
81 |
|
82 section {* Quotient composition lemmas *} |
|
83 |
|
84 lemma list_all2_refl': |
|
85 assumes q: "equivp R" |
|
86 shows "(list_all2 R) r r" |
|
87 by (rule list_all2_refl) (metis equivp_def q) |
|
88 |
|
89 lemma compose_list_refl: |
|
90 assumes q: "equivp R" |
|
91 shows "(list_all2 R OOO op \<approx>) r r" |
|
92 proof |
|
93 have *: "r \<approx> r" by (rule equivp_reflp[OF fset_equivp]) |
|
94 show "list_all2 R r r" by (rule list_all2_refl'[OF q]) |
|
95 with * show "(op \<approx> OO list_all2 R) r r" .. |
|
96 qed |
|
97 |
|
98 lemma map_list_eq_cong: "b \<approx> ba \<Longrightarrow> map f b \<approx> map f ba" |
|
99 unfolding list_eq.simps |
|
100 by (simp only: set_map) |
|
101 |
|
102 lemma quotient_compose_list_g: |
|
103 assumes q: "Quotient R Abs Rep" |
|
104 and e: "equivp R" |
|
105 shows "Quotient ((list_all2 R) OOO (op \<approx>)) |
|
106 (abs_fset \<circ> (map Abs)) ((map Rep) \<circ> rep_fset)" |
|
107 unfolding Quotient_def comp_def |
|
108 proof (intro conjI allI) |
|
109 fix a r s |
|
110 show "abs_fset (map Abs (map Rep (rep_fset a))) = a" |
|
111 by (simp add: abs_o_rep[OF q] Quotient_abs_rep[OF Quotient_fset] map_id) |
|
112 have b: "list_all2 R (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
|
113 by (rule list_all2_refl'[OF e]) |
|
114 have c: "(op \<approx> OO list_all2 R) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
|
115 by (rule, rule equivp_reflp[OF fset_equivp]) (rule b) |
|
116 show "(list_all2 R OOO op \<approx>) (map Rep (rep_fset a)) (map Rep (rep_fset a))" |
|
117 by (rule, rule list_all2_refl'[OF e]) (rule c) |
|
118 show "(list_all2 R OOO op \<approx>) r s = ((list_all2 R OOO op \<approx>) r r \<and> |
|
119 (list_all2 R OOO op \<approx>) s s \<and> abs_fset (map Abs r) = abs_fset (map Abs s))" |
|
120 proof (intro iffI conjI) |
|
121 show "(list_all2 R OOO op \<approx>) r r" by (rule compose_list_refl[OF e]) |
|
122 show "(list_all2 R OOO op \<approx>) s s" by (rule compose_list_refl[OF e]) |
|
123 next |
|
124 assume a: "(list_all2 R OOO op \<approx>) r s" |
|
125 then have b: "map Abs r \<approx> map Abs s" |
|
126 proof (elim pred_compE) |
|
127 fix b ba |
|
128 assume c: "list_all2 R r b" |
|
129 assume d: "b \<approx> ba" |
|
130 assume e: "list_all2 R ba s" |
|
131 have f: "map Abs r = map Abs b" |
|
132 using Quotient_rel[OF list_quotient[OF q]] c by blast |
|
133 have "map Abs ba = map Abs s" |
|
134 using Quotient_rel[OF list_quotient[OF q]] e by blast |
|
135 then have g: "map Abs s = map Abs ba" by simp |
|
136 then show "map Abs r \<approx> map Abs s" using d f map_list_eq_cong by simp |
|
137 qed |
|
138 then show "abs_fset (map Abs r) = abs_fset (map Abs s)" |
|
139 using Quotient_rel[OF Quotient_fset] by blast |
|
140 next |
|
141 assume a: "(list_all2 R OOO op \<approx>) r r \<and> (list_all2 R OOO op \<approx>) s s |
|
142 \<and> abs_fset (map Abs r) = abs_fset (map Abs s)" |
|
143 then have s: "(list_all2 R OOO op \<approx>) s s" by simp |
|
144 have d: "map Abs r \<approx> map Abs s" |
|
145 by (subst Quotient_rel[OF Quotient_fset]) (simp add: a) |
|
146 have b: "map Rep (map Abs r) \<approx> map Rep (map Abs s)" |
|
147 by (rule map_list_eq_cong[OF d]) |
|
148 have y: "list_all2 R (map Rep (map Abs s)) s" |
|
149 by (fact rep_abs_rsp_left[OF list_quotient[OF q], OF list_all2_refl'[OF e, of s]]) |
|
150 have c: "(op \<approx> OO list_all2 R) (map Rep (map Abs r)) s" |
|
151 by (rule pred_compI) (rule b, rule y) |
|
152 have z: "list_all2 R r (map Rep (map Abs r))" |
|
153 by (fact rep_abs_rsp[OF list_quotient[OF q], OF list_all2_refl'[OF e, of r]]) |
|
154 then show "(list_all2 R OOO op \<approx>) r s" |
|
155 using a c pred_compI by simp |
|
156 qed |
|
157 qed |
|
158 |
|
159 lemma quotient_compose_list[quot_thm]: |
|
160 shows "Quotient ((list_all2 op \<approx>) OOO (op \<approx>)) |
|
161 (abs_fset \<circ> (map abs_fset)) ((map rep_fset) \<circ> rep_fset)" |
|
162 by (rule quotient_compose_list_g, rule Quotient_fset, rule list_eq_equivp) |
|
163 |
|
164 |
|
165 |
|
166 subsection {* Respectfulness lemmas for list operations *} |
|
167 |
|
168 lemma list_equiv_rsp [quot_respect]: |
|
169 shows "(op \<approx> ===> op \<approx> ===> op =) op \<approx> op \<approx>" |
|
170 by auto |
|
171 |
|
172 lemma append_rsp [quot_respect]: |
|
173 shows "(op \<approx> ===> op \<approx> ===> op \<approx>) append append" |
|
174 by simp |
|
175 |
|
176 lemma sub_list_rsp [quot_respect]: |
|
177 shows "(op \<approx> ===> op \<approx> ===> op =) sub_list sub_list" |
|
178 by simp |
|
179 |
|
180 lemma memb_rsp [quot_respect]: |
|
181 shows "(op = ===> op \<approx> ===> op =) memb memb" |
|
182 by simp |
|
183 |
|
184 lemma nil_rsp [quot_respect]: |
|
185 shows "(op \<approx>) Nil Nil" |
|
186 by simp |
|
187 |
|
188 lemma cons_rsp [quot_respect]: |
|
189 shows "(op = ===> op \<approx> ===> op \<approx>) Cons Cons" |
|
190 by simp |
|
191 |
|
192 lemma map_rsp [quot_respect]: |
|
193 shows "(op = ===> op \<approx> ===> op \<approx>) map map" |
|
194 by auto |
|
195 |
|
196 lemma set_rsp [quot_respect]: |
|
197 "(op \<approx> ===> op =) set set" |
|
198 by auto |
|
199 |
|
200 lemma inter_list_rsp [quot_respect]: |
|
201 shows "(op \<approx> ===> op \<approx> ===> op \<approx>) inter_list inter_list" |
|
202 by simp |
|
203 |
|
204 lemma removeAll_rsp [quot_respect]: |
|
205 shows "(op = ===> op \<approx> ===> op \<approx>) removeAll removeAll" |
|
206 by simp |
|
207 |
|
208 lemma diff_list_rsp [quot_respect]: |
|
209 shows "(op \<approx> ===> op \<approx> ===> op \<approx>) diff_list diff_list" |
|
210 by simp |
|
211 |
|
212 lemma card_list_rsp [quot_respect]: |
|
213 shows "(op \<approx> ===> op =) card_list card_list" |
|
214 by simp |
|
215 |
|
216 lemma filter_rsp [quot_respect]: |
|
217 shows "(op = ===> op \<approx> ===> op \<approx>) filter filter" |
|
218 by simp |
|
219 |
|
220 lemma memb_commute_fold_list: |
|
221 assumes a: "rsp_fold f" |
|
222 and b: "x \<in> set xs" |
|
223 shows "fold_list f y xs = f x (fold_list f y (removeAll x xs))" |
|
224 using a b by (induct xs) (auto simp add: rsp_fold_def) |
|
225 |
|
226 lemma fold_list_rsp_pre: |
|
227 assumes a: "set xs = set ys" |
|
228 shows "fold_list f z xs = fold_list f z ys" |
|
229 using a |
|
230 apply (induct xs arbitrary: ys) |
|
231 apply (simp) |
|
232 apply (simp (no_asm_use)) |
|
233 apply (rule conjI) |
|
234 apply (rule_tac [!] impI) |
|
235 apply (rule_tac [!] conjI) |
|
236 apply (rule_tac [!] impI) |
|
237 apply (metis insert_absorb) |
|
238 apply (metis List.insert_def List.set.simps(2) List.set_insert fold_list.simps(2)) |
|
239 apply (metis Diff_insert_absorb insertI1 memb_commute_fold_list set_removeAll) |
|
240 apply(drule_tac x="removeAll a ys" in meta_spec) |
|
241 apply(auto) |
|
242 apply(drule meta_mp) |
|
243 apply(blast) |
|
244 by (metis List.set.simps(2) emptyE fold_list.simps(2) in_listsp_conv_set listsp.simps mem_def) |
|
245 |
|
246 lemma fold_list_rsp [quot_respect]: |
|
247 shows "(op = ===> op = ===> op \<approx> ===> op =) fold_list fold_list" |
|
248 unfolding fun_rel_def |
|
249 by(auto intro: fold_list_rsp_pre) |
|
250 |
|
251 lemma concat_rsp_pre: |
|
252 assumes a: "list_all2 op \<approx> x x'" |
|
253 and b: "x' \<approx> y'" |
|
254 and c: "list_all2 op \<approx> y' y" |
|
255 and d: "\<exists>x\<in>set x. xa \<in> set x" |
|
256 shows "\<exists>x\<in>set y. xa \<in> set x" |
|
257 proof - |
|
258 obtain xb where e: "xb \<in> set x" and f: "xa \<in> set xb" using d by auto |
|
259 have "\<exists>y. y \<in> set x' \<and> xb \<approx> y" by (rule list_all2_find_element[OF e a]) |
|
260 then obtain ya where h: "ya \<in> set x'" and i: "xb \<approx> ya" by auto |
|
261 have "ya \<in> set y'" using b h by simp |
|
262 then have "\<exists>yb. yb \<in> set y \<and> ya \<approx> yb" using c by (rule list_all2_find_element) |
|
263 then show ?thesis using f i by auto |
|
264 qed |
|
265 |
|
266 lemma concat_rsp [quot_respect]: |
|
267 shows "(list_all2 op \<approx> OOO op \<approx> ===> op \<approx>) concat concat" |
|
268 proof (rule fun_relI, elim pred_compE) |
|
269 fix a b ba bb |
|
270 assume a: "list_all2 op \<approx> a ba" |
|
271 assume b: "ba \<approx> bb" |
|
272 assume c: "list_all2 op \<approx> bb b" |
|
273 have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" |
|
274 proof |
|
275 fix x |
|
276 show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" |
|
277 proof |
|
278 assume d: "\<exists>xa\<in>set a. x \<in> set xa" |
|
279 show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d]) |
|
280 next |
|
281 assume e: "\<exists>xa\<in>set b. x \<in> set xa" |
|
282 have a': "list_all2 op \<approx> ba a" by (rule list_all2_symp[OF list_eq_equivp, OF a]) |
|
283 have b': "bb \<approx> ba" by (rule equivp_symp[OF list_eq_equivp, OF b]) |
|
284 have c': "list_all2 op \<approx> b bb" by (rule list_all2_symp[OF list_eq_equivp, OF c]) |
|
285 show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e]) |
|
286 qed |
|
287 qed |
|
288 then show "concat a \<approx> concat b" by auto |
|
289 qed |
|
290 |
|
291 |
|
292 |
|
293 section {* Quotient definitions for fsets *} |
|
294 |
|
295 |
|
296 subsection {* Finite sets are a bounded, distributive lattice with minus *} |
|
297 |
|
298 instantiation fset :: (type) "{bounded_lattice_bot, distrib_lattice, minus}" |
|
299 begin |
|
300 |
|
301 quotient_definition |
|
302 "bot :: 'a fset" |
|
303 is "Nil :: 'a list" |
|
304 |
|
305 abbreviation |
|
306 empty_fset ("{||}") |
|
307 where |
|
308 "{||} \<equiv> bot :: 'a fset" |
|
309 |
|
310 quotient_definition |
|
311 "less_eq_fset :: ('a fset \<Rightarrow> 'a fset \<Rightarrow> bool)" |
|
312 is "sub_list :: ('a list \<Rightarrow> 'a list \<Rightarrow> bool)" |
|
313 |
|
314 abbreviation |
|
315 subset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) |
|
316 where |
|
317 "xs |\<subseteq>| ys \<equiv> xs \<le> ys" |
|
318 |
|
319 definition |
|
320 less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" |
|
321 where |
|
322 "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)" |
|
323 |
|
324 abbreviation |
|
325 psubset_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) |
|
326 where |
|
327 "xs |\<subset>| ys \<equiv> xs < ys" |
|
328 |
|
329 quotient_definition |
|
330 "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
331 is "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
332 |
|
333 abbreviation |
|
334 union_fset (infixl "|\<union>|" 65) |
|
335 where |
|
336 "xs |\<union>| ys \<equiv> sup xs (ys::'a fset)" |
|
337 |
|
338 quotient_definition |
|
339 "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
340 is "inter_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
341 |
|
342 abbreviation |
|
343 inter_fset (infixl "|\<inter>|" 65) |
|
344 where |
|
345 "xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)" |
|
346 |
|
347 quotient_definition |
|
348 "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
349 is "diff_list :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
|
350 |
|
351 |
|
352 instance |
|
353 proof |
|
354 fix x y z :: "'a fset" |
|
355 show "x |\<subset>| y \<longleftrightarrow> x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x" |
|
356 unfolding less_fset_def |
|
357 by (descending) (auto) |
|
358 show "x |\<subseteq>| x" by (descending) (simp) |
|
359 show "{||} |\<subseteq>| x" by (descending) (simp) |
|
360 show "x |\<subseteq>| x |\<union>| y" by (descending) (simp) |
|
361 show "y |\<subseteq>| x |\<union>| y" by (descending) (simp) |
|
362 show "x |\<inter>| y |\<subseteq>| x" by (descending) (auto) |
|
363 show "x |\<inter>| y |\<subseteq>| y" by (descending) (auto) |
|
364 show "x |\<union>| (y |\<inter>| z) = x |\<union>| y |\<inter>| (x |\<union>| z)" |
|
365 by (descending) (auto) |
|
366 next |
|
367 fix x y z :: "'a fset" |
|
368 assume a: "x |\<subseteq>| y" |
|
369 assume b: "y |\<subseteq>| z" |
|
370 show "x |\<subseteq>| z" using a b by (descending) (simp) |
|
371 next |
|
372 fix x y :: "'a fset" |
|
373 assume a: "x |\<subseteq>| y" |
|
374 assume b: "y |\<subseteq>| x" |
|
375 show "x = y" using a b by (descending) (auto) |
|
376 next |
|
377 fix x y z :: "'a fset" |
|
378 assume a: "y |\<subseteq>| x" |
|
379 assume b: "z |\<subseteq>| x" |
|
380 show "y |\<union>| z |\<subseteq>| x" using a b by (descending) (simp) |
|
381 next |
|
382 fix x y z :: "'a fset" |
|
383 assume a: "x |\<subseteq>| y" |
|
384 assume b: "x |\<subseteq>| z" |
|
385 show "x |\<subseteq>| y |\<inter>| z" using a b by (descending) (auto) |
|
386 qed |
|
387 |
|
388 end |
|
389 |
|
390 |
|
391 subsection {* Other constants for fsets *} |
|
392 |
|
393 quotient_definition |
|
394 "insert_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
395 is "Cons" |
|
396 |
|
397 syntax |
|
398 "@Insert_fset" :: "args => 'a fset" ("{|(_)|}") |
|
399 |
|
400 translations |
|
401 "{|x, xs|}" == "CONST insert_fset x {|xs|}" |
|
402 "{|x|}" == "CONST insert_fset x {||}" |
|
403 |
|
404 quotient_definition |
|
405 in_fset (infix "|\<in>|" 50) |
|
406 where |
|
407 "in_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> bool" is "memb" |
|
408 |
|
409 abbreviation |
|
410 notin_fset :: "'a \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<notin>|" 50) |
|
411 where |
|
412 "x |\<notin>| S \<equiv> \<not> (x |\<in>| S)" |
|
413 |
|
414 |
|
415 subsection {* Other constants on the Quotient Type *} |
|
416 |
|
417 quotient_definition |
|
418 "card_fset :: 'a fset \<Rightarrow> nat" |
|
419 is card_list |
|
420 |
|
421 quotient_definition |
|
422 "map_fset :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a fset \<Rightarrow> 'b fset" |
|
423 is map |
|
424 |
|
425 quotient_definition |
|
426 "remove_fset :: 'a \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
427 is removeAll |
|
428 |
|
429 quotient_definition |
|
430 "fset :: 'a fset \<Rightarrow> 'a set" |
|
431 is "set" |
|
432 |
|
433 quotient_definition |
|
434 "fold_fset :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a fset \<Rightarrow> 'b" |
|
435 is fold_list |
|
436 |
|
437 quotient_definition |
|
438 "concat_fset :: ('a fset) fset \<Rightarrow> 'a fset" |
|
439 is concat |
|
440 |
|
441 quotient_definition |
|
442 "filter_fset :: ('a \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
|
443 is filter |
|
444 |
|
445 |
|
446 subsection {* Compositional respectfulness and preservation lemmas *} |
|
447 |
|
448 lemma Nil_rsp2 [quot_respect]: |
|
449 shows "(list_all2 op \<approx> OOO op \<approx>) Nil Nil" |
|
450 by (rule compose_list_refl, rule list_eq_equivp) |
|
451 |
|
452 lemma Cons_rsp2 [quot_respect]: |
|
453 shows "(op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) Cons Cons" |
|
454 apply auto |
|
455 apply (rule_tac b="x # b" in pred_compI) |
|
456 apply auto |
|
457 apply (rule_tac b="x # ba" in pred_compI) |
|
458 apply auto |
|
459 done |
|
460 |
|
461 lemma map_prs [quot_preserve]: |
|
462 shows "(abs_fset \<circ> map f) [] = abs_fset []" |
|
463 by simp |
|
464 |
|
465 lemma insert_fset_rsp [quot_preserve]: |
|
466 "(rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) Cons = insert_fset" |
|
467 by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] |
|
468 abs_o_rep[OF Quotient_fset] map_id insert_fset_def) |
|
469 |
|
470 lemma union_fset_rsp [quot_preserve]: |
|
471 "((map rep_fset \<circ> rep_fset) ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) |
|
472 append = union_fset" |
|
473 by (simp add: fun_eq_iff Quotient_abs_rep[OF Quotient_fset] |
|
474 abs_o_rep[OF Quotient_fset] map_id sup_fset_def) |
|
475 |
|
476 lemma list_all2_app_l: |
|
477 assumes a: "reflp R" |
|
478 and b: "list_all2 R l r" |
|
479 shows "list_all2 R (z @ l) (z @ r)" |
|
480 by (induct z) (simp_all add: b rev_iffD1[OF a meta_eq_to_obj_eq[OF reflp_def]]) |
|
481 |
|
482 lemma append_rsp2_pre0: |
|
483 assumes a:"list_all2 op \<approx> x x'" |
|
484 shows "list_all2 op \<approx> (x @ z) (x' @ z)" |
|
485 using a apply (induct x x' rule: list_induct2') |
|
486 by simp_all (rule list_all2_refl'[OF list_eq_equivp]) |
|
487 |
|
488 lemma append_rsp2_pre1: |
|
489 assumes a:"list_all2 op \<approx> x x'" |
|
490 shows "list_all2 op \<approx> (z @ x) (z @ x')" |
|
491 using a apply (induct x x' arbitrary: z rule: list_induct2') |
|
492 apply (rule list_all2_refl'[OF list_eq_equivp]) |
|
493 apply (simp_all del: list_eq.simps) |
|
494 apply (rule list_all2_app_l) |
|
495 apply (simp_all add: reflp_def) |
|
496 done |
|
497 |
|
498 lemma append_rsp2_pre: |
|
499 assumes a:"list_all2 op \<approx> x x'" |
|
500 and b: "list_all2 op \<approx> z z'" |
|
501 shows "list_all2 op \<approx> (x @ z) (x' @ z')" |
|
502 apply (rule list_all2_transp[OF fset_equivp]) |
|
503 apply (rule append_rsp2_pre0) |
|
504 apply (rule a) |
|
505 using b apply (induct z z' rule: list_induct2') |
|
506 apply (simp_all only: append_Nil2) |
|
507 apply (rule list_all2_refl'[OF list_eq_equivp]) |
|
508 apply simp_all |
|
509 apply (rule append_rsp2_pre1) |
|
510 apply simp |
|
511 done |
|
512 |
|
513 lemma append_rsp2 [quot_respect]: |
|
514 "(list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx> ===> list_all2 op \<approx> OOO op \<approx>) append append" |
|
515 proof (intro fun_relI, elim pred_compE) |
|
516 fix x y z w x' z' y' w' :: "'a list list" |
|
517 assume a:"list_all2 op \<approx> x x'" |
|
518 and b: "x' \<approx> y'" |
|
519 and c: "list_all2 op \<approx> y' y" |
|
520 assume aa: "list_all2 op \<approx> z z'" |
|
521 and bb: "z' \<approx> w'" |
|
522 and cc: "list_all2 op \<approx> w' w" |
|
523 have a': "list_all2 op \<approx> (x @ z) (x' @ z')" using a aa append_rsp2_pre by auto |
|
524 have b': "x' @ z' \<approx> y' @ w'" using b bb by simp |
|
525 have c': "list_all2 op \<approx> (y' @ w') (y @ w)" using c cc append_rsp2_pre by auto |
|
526 have d': "(op \<approx> OO list_all2 op \<approx>) (x' @ z') (y @ w)" |
|
527 by (rule pred_compI) (rule b', rule c') |
|
528 show "(list_all2 op \<approx> OOO op \<approx>) (x @ z) (y @ w)" |
|
529 by (rule pred_compI) (rule a', rule d') |
|
530 qed |
|
531 |
|
532 |
|
533 |
|
534 section {* Lifted theorems *} |
|
535 |
|
536 subsection {* fset *} |
|
537 |
|
538 lemma fset_simps [simp]: |
|
539 shows "fset {||} = {}" |
|
540 and "fset (insert_fset x S) = insert x (fset S)" |
|
541 by (descending, simp)+ |
|
542 |
|
543 lemma finite_fset [simp]: |
|
544 shows "finite (fset S)" |
|
545 by (descending) (simp) |
|
546 |
|
547 lemma fset_cong: |
|
548 shows "fset S = fset T \<longleftrightarrow> S = T" |
|
549 by (descending) (simp) |
|
550 |
|
551 lemma filter_fset [simp]: |
|
552 shows "fset (filter_fset P xs) = P \<inter> fset xs" |
|
553 by (descending) (auto simp add: mem_def) |
|
554 |
|
555 lemma remove_fset [simp]: |
|
556 shows "fset (remove_fset x xs) = fset xs - {x}" |
|
557 by (descending) (simp) |
|
558 |
|
559 lemma inter_fset [simp]: |
|
560 shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys" |
|
561 by (descending) (auto) |
|
562 |
|
563 lemma union_fset [simp]: |
|
564 shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys" |
|
565 by (lifting set_append) |
|
566 |
|
567 lemma minus_fset [simp]: |
|
568 shows "fset (xs - ys) = fset xs - fset ys" |
|
569 by (descending) (auto) |
|
570 |
|
571 |
|
572 subsection {* in_fset *} |
|
573 |
|
574 lemma in_fset: |
|
575 shows "x |\<in>| S \<longleftrightarrow> x \<in> fset S" |
|
576 by (descending) (simp) |
|
577 |
|
578 lemma notin_fset: |
|
579 shows "x |\<notin>| S \<longleftrightarrow> x \<notin> fset S" |
|
580 by (simp add: in_fset) |
|
581 |
|
582 lemma notin_empty_fset: |
|
583 shows "x |\<notin>| {||}" |
|
584 by (simp add: in_fset) |
|
585 |
|
586 lemma fset_eq_iff: |
|
587 shows "S = T \<longleftrightarrow> (\<forall>x. (x |\<in>| S) = (x |\<in>| T))" |
|
588 by (descending) (auto) |
|
589 |
|
590 lemma none_in_empty_fset: |
|
591 shows "(\<forall>x. x |\<notin>| S) \<longleftrightarrow> S = {||}" |
|
592 by (descending) (simp) |
|
593 |
|
594 |
|
595 subsection {* insert_fset *} |
|
596 |
|
597 lemma in_insert_fset_iff [simp]: |
|
598 shows "x |\<in>| insert_fset y S \<longleftrightarrow> x = y \<or> x |\<in>| S" |
|
599 by (descending) (simp) |
|
600 |
|
601 lemma |
|
602 shows insert_fsetI1: "x |\<in>| insert_fset x S" |
|
603 and insert_fsetI2: "x |\<in>| S \<Longrightarrow> x |\<in>| insert_fset y S" |
|
604 by simp_all |
|
605 |
|
606 lemma insert_absorb_fset [simp]: |
|
607 shows "x |\<in>| S \<Longrightarrow> insert_fset x S = S" |
|
608 by (descending) (auto) |
|
609 |
|
610 lemma empty_not_insert_fset[simp]: |
|
611 shows "{||} \<noteq> insert_fset x S" |
|
612 and "insert_fset x S \<noteq> {||}" |
|
613 by (descending, simp)+ |
|
614 |
|
615 lemma insert_fset_left_comm: |
|
616 shows "insert_fset x (insert_fset y S) = insert_fset y (insert_fset x S)" |
|
617 by (descending) (auto) |
|
618 |
|
619 lemma insert_fset_left_idem: |
|
620 shows "insert_fset x (insert_fset x S) = insert_fset x S" |
|
621 by (descending) (auto) |
|
622 |
|
623 lemma singleton_fset_eq[simp]: |
|
624 shows "{|x|} = {|y|} \<longleftrightarrow> x = y" |
|
625 by (descending) (auto) |
|
626 |
|
627 lemma in_fset_mdef: |
|
628 shows "x |\<in>| F \<longleftrightarrow> x |\<notin>| (F - {|x|}) \<and> F = insert_fset x (F - {|x|})" |
|
629 by (descending) (auto) |
|
630 |
|
631 |
|
632 subsection {* union_fset *} |
|
633 |
|
634 lemmas [simp] = |
|
635 sup_bot_left[where 'a="'a fset", standard] |
|
636 sup_bot_right[where 'a="'a fset", standard] |
|
637 |
|
638 lemma union_insert_fset [simp]: |
|
639 shows "insert_fset x S |\<union>| T = insert_fset x (S |\<union>| T)" |
|
640 by (lifting append.simps(2)) |
|
641 |
|
642 lemma singleton_union_fset_left: |
|
643 shows "{|a|} |\<union>| S = insert_fset a S" |
|
644 by simp |
|
645 |
|
646 lemma singleton_union_fset_right: |
|
647 shows "S |\<union>| {|a|} = insert_fset a S" |
|
648 by (subst sup.commute) simp |
|
649 |
|
650 lemma in_union_fset: |
|
651 shows "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T" |
|
652 by (descending) (simp) |
|
653 |
|
654 |
|
655 subsection {* minus_fset *} |
|
656 |
|
657 lemma minus_in_fset: |
|
658 shows "x |\<in>| (xs - ys) \<longleftrightarrow> x |\<in>| xs \<and> x |\<notin>| ys" |
|
659 by (descending) (simp) |
|
660 |
|
661 lemma minus_insert_fset: |
|
662 shows "insert_fset x xs - ys = (if x |\<in>| ys then xs - ys else insert_fset x (xs - ys))" |
|
663 by (descending) (auto) |
|
664 |
|
665 lemma minus_insert_in_fset[simp]: |
|
666 shows "x |\<in>| ys \<Longrightarrow> insert_fset x xs - ys = xs - ys" |
|
667 by (simp add: minus_insert_fset) |
|
668 |
|
669 lemma minus_insert_notin_fset[simp]: |
|
670 shows "x |\<notin>| ys \<Longrightarrow> insert_fset x xs - ys = insert_fset x (xs - ys)" |
|
671 by (simp add: minus_insert_fset) |
|
672 |
|
673 lemma in_minus_fset: |
|
674 shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S" |
|
675 unfolding in_fset minus_fset |
|
676 by blast |
|
677 |
|
678 lemma notin_minus_fset: |
|
679 shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S" |
|
680 unfolding in_fset minus_fset |
|
681 by blast |
|
682 |
|
683 |
|
684 subsection {* remove_fset *} |
|
685 |
|
686 lemma in_remove_fset: |
|
687 shows "x |\<in>| remove_fset y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y" |
|
688 by (descending) (simp) |
|
689 |
|
690 lemma notin_remove_fset: |
|
691 shows "x |\<notin>| remove_fset x S" |
|
692 by (descending) (simp) |
|
693 |
|
694 lemma notin_remove_ident_fset: |
|
695 shows "x |\<notin>| S \<Longrightarrow> remove_fset x S = S" |
|
696 by (descending) (simp) |
|
697 |
|
698 lemma remove_fset_cases: |
|
699 shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = insert_fset x (remove_fset x S))" |
|
700 by (descending) (auto simp add: insert_absorb) |
|
701 |
|
702 |
|
703 subsection {* inter_fset *} |
|
704 |
|
705 lemma inter_empty_fset_l: |
|
706 shows "{||} |\<inter>| S = {||}" |
|
707 by simp |
|
708 |
|
709 lemma inter_empty_fset_r: |
|
710 shows "S |\<inter>| {||} = {||}" |
|
711 by simp |
|
712 |
|
713 lemma inter_insert_fset: |
|
714 shows "insert_fset x S |\<inter>| T = (if x |\<in>| T then insert_fset x (S |\<inter>| T) else S |\<inter>| T)" |
|
715 by (descending) (auto) |
|
716 |
|
717 lemma in_inter_fset: |
|
718 shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T" |
|
719 by (descending) (simp) |
|
720 |
|
721 |
|
722 subsection {* subset_fset and psubset_fset *} |
|
723 |
|
724 lemma subset_fset: |
|
725 shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys" |
|
726 by (descending) (simp) |
|
727 |
|
728 lemma psubset_fset: |
|
729 shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys" |
|
730 unfolding less_fset_def |
|
731 by (descending) (auto) |
|
732 |
|
733 lemma subset_insert_fset: |
|
734 shows "(insert_fset x xs) |\<subseteq>| ys \<longleftrightarrow> x |\<in>| ys \<and> xs |\<subseteq>| ys" |
|
735 by (descending) (simp) |
|
736 |
|
737 lemma subset_in_fset: |
|
738 shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)" |
|
739 by (descending) (auto) |
|
740 |
|
741 lemma subset_empty_fset: |
|
742 shows "xs |\<subseteq>| {||} \<longleftrightarrow> xs = {||}" |
|
743 by (descending) (simp) |
|
744 |
|
745 lemma not_psubset_empty_fset: |
|
746 shows "\<not> xs |\<subset>| {||}" |
|
747 by (metis fset_simps(1) psubset_fset not_psubset_empty) |
|
748 |
|
749 |
|
750 subsection {* map_fset *} |
|
751 |
|
752 lemma map_fset_simps [simp]: |
|
753 shows "map_fset f {||} = {||}" |
|
754 and "map_fset f (insert_fset x S) = insert_fset (f x) (map_fset f S)" |
|
755 by (descending, simp)+ |
|
756 |
|
757 lemma map_fset_image [simp]: |
|
758 shows "fset (map_fset f S) = f ` (fset S)" |
|
759 by (descending) (simp) |
|
760 |
|
761 lemma inj_map_fset_cong: |
|
762 shows "inj f \<Longrightarrow> map_fset f S = map_fset f T \<longleftrightarrow> S = T" |
|
763 by (descending) (metis inj_vimage_image_eq list_eq.simps set_map) |
|
764 |
|
765 lemma map_union_fset: |
|
766 shows "map_fset f (S |\<union>| T) = map_fset f S |\<union>| map_fset f T" |
|
767 by (descending) (simp) |
|
768 |
|
769 |
|
770 subsection {* card_fset *} |
|
771 |
|
772 lemma card_fset: |
|
773 shows "card_fset xs = card (fset xs)" |
|
774 by (descending) (simp) |
|
775 |
|
776 lemma card_insert_fset_iff [simp]: |
|
777 shows "card_fset (insert_fset x S) = (if x |\<in>| S then card_fset S else Suc (card_fset S))" |
|
778 by (descending) (simp add: insert_absorb) |
|
779 |
|
780 lemma card_fset_0[simp]: |
|
781 shows "card_fset S = 0 \<longleftrightarrow> S = {||}" |
|
782 by (descending) (simp) |
|
783 |
|
784 lemma card_empty_fset[simp]: |
|
785 shows "card_fset {||} = 0" |
|
786 by (simp add: card_fset) |
|
787 |
|
788 lemma card_fset_1: |
|
789 shows "card_fset S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})" |
|
790 by (descending) (auto simp add: card_Suc_eq) |
|
791 |
|
792 lemma card_fset_gt_0: |
|
793 shows "x \<in> fset S \<Longrightarrow> 0 < card_fset S" |
|
794 by (descending) (auto simp add: card_gt_0_iff) |
|
795 |
|
796 lemma card_notin_fset: |
|
797 shows "(x |\<notin>| S) = (card_fset (insert_fset x S) = Suc (card_fset S))" |
|
798 by simp |
|
799 |
|
800 lemma card_fset_Suc: |
|
801 shows "card_fset S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = insert_fset x T \<and> card_fset T = n" |
|
802 apply(descending) |
|
803 apply(auto dest!: card_eq_SucD) |
|
804 by (metis Diff_insert_absorb set_removeAll) |
|
805 |
|
806 lemma card_remove_fset_iff [simp]: |
|
807 shows "card_fset (remove_fset y S) = (if y |\<in>| S then card_fset S - 1 else card_fset S)" |
|
808 by (descending) (simp) |
|
809 |
|
810 lemma card_Suc_exists_in_fset: |
|
811 shows "card_fset S = Suc n \<Longrightarrow> \<exists>a. a |\<in>| S" |
|
812 by (drule card_fset_Suc) (auto) |
|
813 |
|
814 lemma in_card_fset_not_0: |
|
815 shows "a |\<in>| A \<Longrightarrow> card_fset A \<noteq> 0" |
|
816 by (descending) (auto) |
|
817 |
|
818 lemma card_fset_mono: |
|
819 shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset xs \<le> card_fset ys" |
|
820 unfolding card_fset psubset_fset |
|
821 by (simp add: card_mono subset_fset) |
|
822 |
|
823 lemma card_subset_fset_eq: |
|
824 shows "xs |\<subseteq>| ys \<Longrightarrow> card_fset ys \<le> card_fset xs \<Longrightarrow> xs = ys" |
|
825 unfolding card_fset subset_fset |
|
826 by (auto dest: card_seteq[OF finite_fset] simp add: fset_cong) |
|
827 |
|
828 lemma psubset_card_fset_mono: |
|
829 shows "xs |\<subset>| ys \<Longrightarrow> card_fset xs < card_fset ys" |
|
830 unfolding card_fset subset_fset |
|
831 by (metis finite_fset psubset_fset psubset_card_mono) |
|
832 |
|
833 lemma card_union_inter_fset: |
|
834 shows "card_fset xs + card_fset ys = card_fset (xs |\<union>| ys) + card_fset (xs |\<inter>| ys)" |
|
835 unfolding card_fset union_fset inter_fset |
|
836 by (rule card_Un_Int[OF finite_fset finite_fset]) |
|
837 |
|
838 lemma card_union_disjoint_fset: |
|
839 shows "xs |\<inter>| ys = {||} \<Longrightarrow> card_fset (xs |\<union>| ys) = card_fset xs + card_fset ys" |
|
840 unfolding card_fset union_fset |
|
841 apply (rule card_Un_disjoint[OF finite_fset finite_fset]) |
|
842 by (metis inter_fset fset_simps(1)) |
|
843 |
|
844 lemma card_remove_fset_less1: |
|
845 shows "x |\<in>| xs \<Longrightarrow> card_fset (remove_fset x xs) < card_fset xs" |
|
846 unfolding card_fset in_fset remove_fset |
|
847 by (rule card_Diff1_less[OF finite_fset]) |
|
848 |
|
849 lemma card_remove_fset_less2: |
|
850 shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> card_fset (remove_fset y (remove_fset x xs)) < card_fset xs" |
|
851 unfolding card_fset remove_fset in_fset |
|
852 by (rule card_Diff2_less[OF finite_fset]) |
|
853 |
|
854 lemma card_remove_fset_le1: |
|
855 shows "card_fset (remove_fset x xs) \<le> card_fset xs" |
|
856 unfolding remove_fset card_fset |
|
857 by (rule card_Diff1_le[OF finite_fset]) |
|
858 |
|
859 lemma card_psubset_fset: |
|
860 shows "ys |\<subseteq>| xs \<Longrightarrow> card_fset ys < card_fset xs \<Longrightarrow> ys |\<subset>| xs" |
|
861 unfolding card_fset psubset_fset subset_fset |
|
862 by (rule card_psubset[OF finite_fset]) |
|
863 |
|
864 lemma card_map_fset_le: |
|
865 shows "card_fset (map_fset f xs) \<le> card_fset xs" |
|
866 unfolding card_fset map_fset_image |
|
867 by (rule card_image_le[OF finite_fset]) |
|
868 |
|
869 lemma card_minus_insert_fset[simp]: |
|
870 assumes "a |\<in>| A" and "a |\<notin>| B" |
|
871 shows "card_fset (A - insert_fset a B) = card_fset (A - B) - 1" |
|
872 using assms |
|
873 unfolding in_fset card_fset minus_fset |
|
874 by (simp add: card_Diff_insert[OF finite_fset]) |
|
875 |
|
876 lemma card_minus_subset_fset: |
|
877 assumes "B |\<subseteq>| A" |
|
878 shows "card_fset (A - B) = card_fset A - card_fset B" |
|
879 using assms |
|
880 unfolding subset_fset card_fset minus_fset |
|
881 by (rule card_Diff_subset[OF finite_fset]) |
|
882 |
|
883 lemma card_minus_fset: |
|
884 shows "card_fset (A - B) = card_fset A - card_fset (A |\<inter>| B)" |
|
885 unfolding inter_fset card_fset minus_fset |
|
886 by (rule card_Diff_subset_Int) (simp) |
|
887 |
|
888 |
|
889 subsection {* concat_fset *} |
|
890 |
|
891 lemma concat_empty_fset [simp]: |
|
892 shows "concat_fset {||} = {||}" |
|
893 by (lifting concat.simps(1)) |
|
894 |
|
895 lemma concat_insert_fset [simp]: |
|
896 shows "concat_fset (insert_fset x S) = x |\<union>| concat_fset S" |
|
897 by (lifting concat.simps(2)) |
|
898 |
|
899 lemma concat_inter_fset [simp]: |
|
900 shows "concat_fset (xs |\<union>| ys) = concat_fset xs |\<union>| concat_fset ys" |
|
901 by (lifting concat_append) |
|
902 |
|
903 |
|
904 subsection {* filter_fset *} |
|
905 |
|
906 lemma subset_filter_fset: |
|
907 shows "filter_fset P xs |\<subseteq>| filter_fset Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)" |
|
908 by (descending) (auto) |
|
909 |
|
910 lemma eq_filter_fset: |
|
911 shows "(filter_fset P xs = filter_fset Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)" |
|
912 by (descending) (auto) |
|
913 |
|
914 lemma psubset_filter_fset: |
|
915 shows "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> |
|
916 filter_fset P xs |\<subset>| filter_fset Q xs" |
|
917 unfolding less_fset_def by (auto simp add: subset_filter_fset eq_filter_fset) |
|
918 |
|
919 |
|
920 subsection {* fold_fset *} |
|
921 |
|
922 lemma fold_empty_fset: |
|
923 shows "fold_fset f z {||} = z" |
|
924 by (descending) (simp) |
|
925 |
|
926 lemma fold_insert_fset: "fold_fset f z (insert_fset a A) = |
|
927 (if rsp_fold f then if a |\<in>| A then fold_fset f z A else f a (fold_fset f z A) else z)" |
|
928 by (descending) (simp) |
|
929 |
|
930 lemma in_commute_fold_fset: |
|
931 "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> fold_fset f z b = f h (fold_fset f z (remove_fset h b))" |
|
932 by (descending) (simp add: memb_commute_fold_list) |
|
933 |
|
934 |
|
935 subsection {* Choice in fsets *} |
|
936 |
|
937 lemma fset_choice: |
|
938 assumes a: "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y)" |
|
939 shows "\<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)" |
|
940 using a |
|
941 apply(descending) |
|
942 using finite_set_choice |
|
943 by (auto simp add: Ball_def) |
|
944 |
|
945 |
|
946 section {* Induction and Cases rules for fsets *} |
|
947 |
|
948 lemma fset_exhaust [case_names empty_fset insert_fset, cases type: fset]: |
|
949 assumes empty_fset_case: "S = {||} \<Longrightarrow> P" |
|
950 and insert_fset_case: "\<And>x S'. S = insert_fset x S' \<Longrightarrow> P" |
|
951 shows "P" |
|
952 using assms by (lifting list.exhaust) |
|
953 |
|
954 lemma fset_induct [case_names empty_fset insert_fset]: |
|
955 assumes empty_fset_case: "P {||}" |
|
956 and insert_fset_case: "\<And>x S. P S \<Longrightarrow> P (insert_fset x S)" |
|
957 shows "P S" |
|
958 using assms |
|
959 by (descending) (blast intro: list.induct) |
|
960 |
|
961 lemma fset_induct_stronger [case_names empty_fset insert_fset, induct type: fset]: |
|
962 assumes empty_fset_case: "P {||}" |
|
963 and insert_fset_case: "\<And>x S. \<lbrakk>x |\<notin>| S; P S\<rbrakk> \<Longrightarrow> P (insert_fset x S)" |
|
964 shows "P S" |
|
965 proof(induct S rule: fset_induct) |
|
966 case empty_fset |
|
967 show "P {||}" using empty_fset_case by simp |
|
968 next |
|
969 case (insert_fset x S) |
|
970 have "P S" by fact |
|
971 then show "P (insert_fset x S)" using insert_fset_case |
|
972 by (cases "x |\<in>| S") (simp_all) |
|
973 qed |
|
974 |
|
975 lemma fset_card_induct: |
|
976 assumes empty_fset_case: "P {||}" |
|
977 and card_fset_Suc_case: "\<And>S T. Suc (card_fset S) = (card_fset T) \<Longrightarrow> P S \<Longrightarrow> P T" |
|
978 shows "P S" |
|
979 proof (induct S) |
|
980 case empty_fset |
|
981 show "P {||}" by (rule empty_fset_case) |
|
982 next |
|
983 case (insert_fset x S) |
|
984 have h: "P S" by fact |
|
985 have "x |\<notin>| S" by fact |
|
986 then have "Suc (card_fset S) = card_fset (insert_fset x S)" |
|
987 using card_fset_Suc by auto |
|
988 then show "P (insert_fset x S)" |
|
989 using h card_fset_Suc_case by simp |
|
990 qed |
|
991 |
|
992 lemma fset_raw_strong_cases: |
|
993 obtains "xs = []" |
|
994 | x ys where "\<not> memb x ys" and "xs \<approx> x # ys" |
|
995 proof (induct xs arbitrary: x ys) |
|
996 case Nil |
|
997 then show thesis by simp |
|
998 next |
|
999 case (Cons a xs) |
|
1000 have a: "\<lbrakk>xs = [] \<Longrightarrow> thesis; \<And>x ys. \<lbrakk>\<not> memb x ys; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis\<rbrakk> \<Longrightarrow> thesis" by fact |
|
1001 have b: "\<And>x' ys'. \<lbrakk>\<not> memb x' ys'; a # xs \<approx> x' # ys'\<rbrakk> \<Longrightarrow> thesis" by fact |
|
1002 have c: "xs = [] \<Longrightarrow> thesis" using b |
|
1003 apply(simp) |
|
1004 by (metis List.set.simps(1) emptyE empty_subsetI) |
|
1005 have "\<And>x ys. \<lbrakk>\<not> memb x ys; xs \<approx> x # ys\<rbrakk> \<Longrightarrow> thesis" |
|
1006 proof - |
|
1007 fix x :: 'a |
|
1008 fix ys :: "'a list" |
|
1009 assume d:"\<not> memb x ys" |
|
1010 assume e:"xs \<approx> x # ys" |
|
1011 show thesis |
|
1012 proof (cases "x = a") |
|
1013 assume h: "x = a" |
|
1014 then have f: "\<not> memb a ys" using d by simp |
|
1015 have g: "a # xs \<approx> a # ys" using e h by auto |
|
1016 show thesis using b f g by simp |
|
1017 next |
|
1018 assume h: "x \<noteq> a" |
|
1019 then have f: "\<not> memb x (a # ys)" using d by auto |
|
1020 have g: "a # xs \<approx> x # (a # ys)" using e h by auto |
|
1021 show thesis using b f g by (simp del: memb_def) |
|
1022 qed |
|
1023 qed |
|
1024 then show thesis using a c by blast |
|
1025 qed |
|
1026 |
|
1027 |
|
1028 lemma fset_strong_cases: |
|
1029 obtains "xs = {||}" |
|
1030 | x ys where "x |\<notin>| ys" and "xs = insert_fset x ys" |
|
1031 by (lifting fset_raw_strong_cases) |
|
1032 |
|
1033 |
|
1034 lemma fset_induct2: |
|
1035 "P {||} {||} \<Longrightarrow> |
|
1036 (\<And>x xs. x |\<notin>| xs \<Longrightarrow> P (insert_fset x xs) {||}) \<Longrightarrow> |
|
1037 (\<And>y ys. y |\<notin>| ys \<Longrightarrow> P {||} (insert_fset y ys)) \<Longrightarrow> |
|
1038 (\<And>x xs y ys. \<lbrakk>P xs ys; x |\<notin>| xs; y |\<notin>| ys\<rbrakk> \<Longrightarrow> P (insert_fset x xs) (insert_fset y ys)) \<Longrightarrow> |
|
1039 P xsa ysa" |
|
1040 apply (induct xsa arbitrary: ysa) |
|
1041 apply (induct_tac x rule: fset_induct_stronger) |
|
1042 apply simp_all |
|
1043 apply (induct_tac xa rule: fset_induct_stronger) |
|
1044 apply simp_all |
|
1045 done |
|
1046 |
|
1047 |
|
1048 |
|
1049 subsection {* alternate formulation with a different decomposition principle |
|
1050 and a proof of equivalence *} |
|
1051 |
|
1052 inductive |
|
1053 list_eq2 ("_ \<approx>2 _") |
|
1054 where |
|
1055 "(a # b # xs) \<approx>2 (b # a # xs)" |
|
1056 | "[] \<approx>2 []" |
|
1057 | "xs \<approx>2 ys \<Longrightarrow> ys \<approx>2 xs" |
|
1058 | "(a # a # xs) \<approx>2 (a # xs)" |
|
1059 | "xs \<approx>2 ys \<Longrightarrow> (a # xs) \<approx>2 (a # ys)" |
|
1060 | "\<lbrakk>xs1 \<approx>2 xs2; xs2 \<approx>2 xs3\<rbrakk> \<Longrightarrow> xs1 \<approx>2 xs3" |
|
1061 |
|
1062 lemma list_eq2_refl: |
|
1063 shows "xs \<approx>2 xs" |
|
1064 by (induct xs) (auto intro: list_eq2.intros) |
|
1065 |
|
1066 lemma cons_delete_list_eq2: |
|
1067 shows "(a # (removeAll a A)) \<approx>2 (if memb a A then A else a # A)" |
|
1068 apply (induct A) |
|
1069 apply (simp add: list_eq2_refl) |
|
1070 apply (case_tac "memb a (aa # A)") |
|
1071 apply (simp_all) |
|
1072 apply (case_tac [!] "a = aa") |
|
1073 apply (simp_all) |
|
1074 apply (case_tac "memb a A") |
|
1075 apply (auto)[2] |
|
1076 apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6)) |
|
1077 apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6)) |
|
1078 apply (auto simp add: list_eq2_refl memb_def) |
|
1079 done |
|
1080 |
|
1081 lemma memb_delete_list_eq2: |
|
1082 assumes a: "memb e r" |
|
1083 shows "(e # removeAll e r) \<approx>2 r" |
|
1084 using a cons_delete_list_eq2[of e r] |
|
1085 by simp |
|
1086 |
|
1087 lemma list_eq2_equiv: |
|
1088 "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)" |
|
1089 proof |
|
1090 show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto |
|
1091 next |
|
1092 { |
|
1093 fix n |
|
1094 assume a: "card_list l = n" and b: "l \<approx> r" |
|
1095 have "l \<approx>2 r" |
|
1096 using a b |
|
1097 proof (induct n arbitrary: l r) |
|
1098 case 0 |
|
1099 have "card_list l = 0" by fact |
|
1100 then have "\<forall>x. \<not> memb x l" by auto |
|
1101 then have z: "l = []" by auto |
|
1102 then have "r = []" using `l \<approx> r` by simp |
|
1103 then show ?case using z list_eq2_refl by simp |
|
1104 next |
|
1105 case (Suc m) |
|
1106 have b: "l \<approx> r" by fact |
|
1107 have d: "card_list l = Suc m" by fact |
|
1108 then have "\<exists>a. memb a l" |
|
1109 apply(simp) |
|
1110 apply(drule card_eq_SucD) |
|
1111 apply(blast) |
|
1112 done |
|
1113 then obtain a where e: "memb a l" by auto |
|
1114 then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b |
|
1115 by auto |
|
1116 have f: "card_list (removeAll a l) = m" using e d by (simp) |
|
1117 have g: "removeAll a l \<approx> removeAll a r" using removeAll_rsp b by simp |
|
1118 have "(removeAll a l) \<approx>2 (removeAll a r)" by (rule Suc.hyps[OF f g]) |
|
1119 then have h: "(a # removeAll a l) \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(5)) |
|
1120 have i: "l \<approx>2 (a # removeAll a l)" |
|
1121 by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]]) |
|
1122 have "l \<approx>2 (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h]) |
|
1123 then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp |
|
1124 qed |
|
1125 } |
|
1126 then show "l \<approx> r \<Longrightarrow> l \<approx>2 r" by blast |
|
1127 qed |
|
1128 |
|
1129 |
|
1130 (* We cannot write it as "assumes .. shows" since Isabelle changes |
|
1131 the quantifiers to schematic variables and reintroduces them in |
|
1132 a different order *) |
|
1133 lemma fset_eq_cases: |
|
1134 "\<lbrakk>a1 = a2; |
|
1135 \<And>a b xs. \<lbrakk>a1 = insert_fset a (insert_fset b xs); a2 = insert_fset b (insert_fset a xs)\<rbrakk> \<Longrightarrow> P; |
|
1136 \<lbrakk>a1 = {||}; a2 = {||}\<rbrakk> \<Longrightarrow> P; \<And>xs ys. \<lbrakk>a1 = ys; a2 = xs; xs = ys\<rbrakk> \<Longrightarrow> P; |
|
1137 \<And>a xs. \<lbrakk>a1 = insert_fset a (insert_fset a xs); a2 = insert_fset a xs\<rbrakk> \<Longrightarrow> P; |
|
1138 \<And>xs ys a. \<lbrakk>a1 = insert_fset a xs; a2 = insert_fset a ys; xs = ys\<rbrakk> \<Longrightarrow> P; |
|
1139 \<And>xs1 xs2 xs3. \<lbrakk>a1 = xs1; a2 = xs3; xs1 = xs2; xs2 = xs3\<rbrakk> \<Longrightarrow> P\<rbrakk> |
|
1140 \<Longrightarrow> P" |
|
1141 by (lifting list_eq2.cases[simplified list_eq2_equiv[symmetric]]) |
|
1142 |
|
1143 lemma fset_eq_induct: |
|
1144 assumes "x1 = x2" |
|
1145 and "\<And>a b xs. P (insert_fset a (insert_fset b xs)) (insert_fset b (insert_fset a xs))" |
|
1146 and "P {||} {||}" |
|
1147 and "\<And>xs ys. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P ys xs" |
|
1148 and "\<And>a xs. P (insert_fset a (insert_fset a xs)) (insert_fset a xs)" |
|
1149 and "\<And>xs ys a. \<lbrakk>xs = ys; P xs ys\<rbrakk> \<Longrightarrow> P (insert_fset a xs) (insert_fset a ys)" |
|
1150 and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3" |
|
1151 shows "P x1 x2" |
|
1152 using assms |
|
1153 by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]]) |
|
1154 |
|
1155 ML {* |
|
1156 fun dest_fsetT (Type (@{type_name fset}, [T])) = T |
|
1157 | dest_fsetT T = raise TYPE ("dest_fsetT: fset type expected", [T], []); |
|
1158 *} |
|
1159 |
|
1160 no_notation |
|
1161 list_eq (infix "\<approx>" 50) and |
|
1162 list_eq2 (infix "\<approx>2" 50) |
|
1163 |
|
1164 end |
|