207 |
201 |
208 lemma memb_cons_iff: |
202 lemma memb_cons_iff: |
209 shows "memb x (y # xs) = (x = y \<or> memb x xs)" |
203 shows "memb x (y # xs) = (x = y \<or> memb x xs)" |
210 by (induct xs) (auto simp add: memb_def) |
204 by (induct xs) (auto simp add: memb_def) |
211 |
205 |
212 lemma memb_finter_raw: |
206 lemma set_finter_raw[simp]: |
213 "memb x (finter_raw xs ys) \<longleftrightarrow> memb x xs \<and> memb x ys" |
207 "set (finter_raw xs ys) = set xs \<inter> set ys" |
214 by (induct xs) (auto simp add: not_memb_nil memb_cons_iff) |
208 by (induct xs) (auto simp add: memb_def) |
215 |
209 |
216 lemma [quot_respect]: |
210 lemma [quot_respect]: |
217 "(op \<approx> ===> op \<approx> ===> op \<approx>) finter_raw finter_raw" |
211 shows "(op \<approx> ===> op \<approx> ===> op \<approx>) finter_raw finter_raw" |
218 by (simp add: memb_def[symmetric] memb_finter_raw) |
212 by (simp) |
219 |
213 |
220 lemma memb_delete_raw: |
214 (* |
221 "memb x (delete_raw xs y) = (memb x xs \<and> x \<noteq> y)" |
215 lemma memb_removeAll: |
|
216 "memb x (removeAll y xs) \<longleftrightarrow> memb x xs \<and> x \<noteq> y" |
222 by (induct xs arbitrary: x y) (auto simp add: memb_def) |
217 by (induct xs arbitrary: x y) (auto simp add: memb_def) |
|
218 *) |
223 |
219 |
224 lemma [quot_respect]: |
220 lemma [quot_respect]: |
225 "(op \<approx> ===> op = ===> op \<approx>) delete_raw delete_raw" |
221 shows "(op = ===> op \<approx> ===> op \<approx>) removeAll removeAll" |
226 by (simp add: memb_def[symmetric] memb_delete_raw) |
222 by (simp) |
227 |
223 |
228 lemma fminus_raw_memb: "memb x (fminus_raw xs ys) = (memb x xs \<and> \<not> memb x ys)" |
224 lemma set_fminus_raw[simp]: |
229 by (induct ys arbitrary: xs) |
225 "set (fminus_raw xs ys) = (set xs - set ys)" |
230 (simp_all add: not_memb_nil memb_delete_raw memb_cons_iff) |
226 by (induct ys arbitrary: xs) (auto) |
231 |
227 |
232 lemma [quot_respect]: |
228 lemma [quot_respect]: |
233 "(op \<approx> ===> op \<approx> ===> op \<approx>) fminus_raw fminus_raw" |
229 shows "(op \<approx> ===> op \<approx> ===> op \<approx>) fminus_raw fminus_raw" |
234 by (simp add: memb_def[symmetric] fminus_raw_memb) |
230 by simp |
235 |
|
236 lemma fcard_raw_gt_0: |
|
237 assumes a: "x \<in> set xs" |
|
238 shows "0 < fcard_raw xs" |
|
239 using a by (induct xs) (auto simp add: memb_def) |
|
240 |
|
241 lemma fcard_raw_delete_one: |
|
242 shows "fcard_raw ([x \<leftarrow> xs. x \<noteq> y]) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)" |
|
243 by (induct xs) (auto dest: fcard_raw_gt_0 simp add: memb_def) |
|
244 |
|
245 lemma fcard_raw_rsp_aux: |
|
246 assumes a: "xs \<approx> ys" |
|
247 shows "fcard_raw xs = fcard_raw ys" |
|
248 using a |
|
249 proof (induct xs arbitrary: ys) |
|
250 case Nil |
|
251 show ?case using Nil.prems by simp |
|
252 next |
|
253 case (Cons a xs) |
|
254 have a: "a # xs \<approx> ys" by fact |
|
255 have b: "\<And>ys. xs \<approx> ys \<Longrightarrow> fcard_raw xs = fcard_raw ys" by fact |
|
256 show ?case proof (cases "a \<in> set xs") |
|
257 assume c: "a \<in> set xs" |
|
258 have "\<forall>x. (x \<in> set xs) = (x \<in> set ys)" |
|
259 proof (intro allI iffI) |
|
260 fix x |
|
261 assume "x \<in> set xs" |
|
262 then show "x \<in> set ys" using a by auto |
|
263 next |
|
264 fix x |
|
265 assume d: "x \<in> set ys" |
|
266 have e: "(x \<in> set (a # xs)) = (x \<in> set ys)" using a by simp |
|
267 show "x \<in> set xs" using c d e unfolding list_eq.simps by simp blast |
|
268 qed |
|
269 then show ?thesis using b c by (simp add: memb_def) |
|
270 next |
|
271 assume c: "a \<notin> set xs" |
|
272 have d: "xs \<approx> [x\<leftarrow>ys . x \<noteq> a] \<Longrightarrow> fcard_raw xs = fcard_raw [x\<leftarrow>ys . x \<noteq> a]" using b by simp |
|
273 have "Suc (fcard_raw xs) = fcard_raw ys" |
|
274 proof (cases "a \<in> set ys") |
|
275 assume e: "a \<in> set ys" |
|
276 have f: "\<forall>x. (x \<in> set xs) = (x \<in> set ys \<and> x \<noteq> a)" using a c |
|
277 by (auto simp add: fcard_raw_delete_one) |
|
278 have "fcard_raw ys = Suc (fcard_raw ys - 1)" by (rule Suc_pred'[OF fcard_raw_gt_0]) (rule e) |
|
279 then show ?thesis using d e f by (simp_all add: fcard_raw_delete_one memb_def) |
|
280 next |
|
281 case False then show ?thesis using a c d by auto |
|
282 qed |
|
283 then show ?thesis using a c d by (simp add: memb_def) |
|
284 qed |
|
285 qed |
|
286 |
231 |
287 lemma fcard_raw_rsp[quot_respect]: |
232 lemma fcard_raw_rsp[quot_respect]: |
288 shows "(op \<approx> ===> op =) fcard_raw fcard_raw" |
233 shows "(op \<approx> ===> op =) fcard_raw fcard_raw" |
289 by (simp add: fcard_raw_rsp_aux) |
234 by (simp add: fcard_raw_def) |
290 |
235 |
291 lemma memb_absorb: |
236 lemma memb_absorb: |
292 shows "memb x xs \<Longrightarrow> x # xs \<approx> xs" |
237 shows "memb x xs \<Longrightarrow> x # xs \<approx> xs" |
293 by (induct xs) (auto simp add: memb_def) |
238 by (induct xs) (auto simp add: memb_def) |
294 |
239 |
295 lemma none_memb_nil: |
240 lemma none_memb_nil: |
296 "(\<forall>x. \<not> memb x xs) = (xs \<approx> [])" |
241 "(\<forall>x. \<not> memb x xs) = (xs \<approx> [])" |
297 by (simp add: memb_def) |
242 by (simp add: memb_def) |
298 |
243 |
299 lemma not_memb_delete_raw_ident: |
244 lemma not_memb_removeAll_ident: |
300 shows "\<not> memb x xs \<Longrightarrow> delete_raw xs x = xs" |
245 shows "\<not> memb x xs \<Longrightarrow> removeAll x xs = xs" |
301 by (induct xs) (auto simp add: memb_def) |
246 by (induct xs) (auto simp add: memb_def) |
302 |
247 |
303 lemma memb_commute_ffold_raw: |
248 lemma memb_commute_ffold_raw: |
304 "rsp_fold f \<Longrightarrow> memb h b \<Longrightarrow> ffold_raw f z b = f h (ffold_raw f z (delete_raw b h))" |
249 "rsp_fold f \<Longrightarrow> h \<in> set b \<Longrightarrow> ffold_raw f z b = f h (ffold_raw f z (removeAll h b))" |
305 apply (induct b) |
250 apply (induct b) |
306 apply (simp_all add: not_memb_nil) |
251 apply (auto simp add: rsp_fold_def) |
307 apply (auto) |
|
308 apply (simp_all add: memb_delete_raw not_memb_delete_raw_ident rsp_fold_def memb_cons_iff) |
|
309 done |
252 done |
310 |
253 |
311 lemma ffold_raw_rsp_pre: |
254 lemma ffold_raw_rsp_pre: |
312 "\<forall>e. memb e a = memb e b \<Longrightarrow> ffold_raw f z a = ffold_raw f z b" |
255 "set a = set b \<Longrightarrow> ffold_raw f z a = ffold_raw f z b" |
313 apply (induct a arbitrary: b) |
256 apply (induct a arbitrary: b) |
314 apply (simp add: memb_absorb memb_def none_memb_nil) |
|
315 apply (simp) |
257 apply (simp) |
|
258 apply (simp (no_asm_use)) |
316 apply (rule conjI) |
259 apply (rule conjI) |
317 apply (rule_tac [!] impI) |
260 apply (rule_tac [!] impI) |
318 apply (rule_tac [!] conjI) |
261 apply (rule_tac [!] conjI) |
319 apply (rule_tac [!] impI) |
262 apply (rule_tac [!] impI) |
320 apply (subgoal_tac "\<forall>e. memb e a2 = memb e b") |
263 apply (metis insert_absorb) |
321 apply (simp) |
264 apply (metis List.insert_def List.set.simps(2) List.set_insert ffold_raw.simps(2)) |
322 apply (simp add: memb_cons_iff memb_def) |
265 apply (metis Diff_insert_absorb insertI1 memb_commute_ffold_raw set_removeAll) |
323 apply (auto)[1] |
266 apply(drule_tac x="removeAll a1 b" in meta_spec) |
324 apply (drule_tac x="e" in spec) |
267 apply(auto) |
325 apply (blast) |
268 apply(drule meta_mp) |
326 apply (case_tac b) |
269 apply(blast) |
327 apply (simp_all) |
270 by (metis List.set.simps(2) emptyE ffold_raw.simps(2) in_listsp_conv_set listsp.simps mem_def) |
328 apply (subgoal_tac "ffold_raw f z b = f a1 (ffold_raw f z (delete_raw b a1))") |
271 |
329 apply (simp only:) |
272 lemma [quot_respect]: |
330 apply (rule_tac f="f a1" in arg_cong) |
273 shows "(op = ===> op = ===> op \<approx> ===> op =) ffold_raw ffold_raw" |
331 apply (subgoal_tac "\<forall>e. memb e a2 = memb e (delete_raw b a1)") |
274 unfolding fun_rel_def |
332 apply (simp) |
275 apply(auto intro: ffold_raw_rsp_pre) |
333 apply (simp add: memb_delete_raw) |
|
334 apply (auto simp add: memb_cons_iff)[1] |
|
335 apply (erule memb_commute_ffold_raw) |
|
336 apply (drule_tac x="a1" in spec) |
|
337 apply (simp add: memb_cons_iff) |
|
338 apply (simp add: memb_cons_iff) |
|
339 apply (case_tac b) |
|
340 apply (simp_all) |
|
341 done |
276 done |
342 |
|
343 lemma [quot_respect]: |
|
344 "(op = ===> op = ===> op \<approx> ===> op =) ffold_raw ffold_raw" |
|
345 by (simp add: memb_def[symmetric] ffold_raw_rsp_pre) |
|
346 |
277 |
347 lemma concat_rsp_pre: |
278 lemma concat_rsp_pre: |
348 assumes a: "list_all2 op \<approx> x x'" |
279 assumes a: "list_all2 op \<approx> x x'" |
349 and b: "x' \<approx> y'" |
280 and b: "x' \<approx> y'" |
350 and c: "list_all2 op \<approx> y' y" |
281 and c: "list_all2 op \<approx> y' y" |
364 proof (rule fun_relI, elim pred_compE) |
295 proof (rule fun_relI, elim pred_compE) |
365 fix a b ba bb |
296 fix a b ba bb |
366 assume a: "list_all2 op \<approx> a ba" |
297 assume a: "list_all2 op \<approx> a ba" |
367 assume b: "ba \<approx> bb" |
298 assume b: "ba \<approx> bb" |
368 assume c: "list_all2 op \<approx> bb b" |
299 assume c: "list_all2 op \<approx> bb b" |
369 have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof |
300 have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" |
|
301 proof |
370 fix x |
302 fix x |
371 show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof |
303 show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" |
|
304 proof |
372 assume d: "\<exists>xa\<in>set a. x \<in> set xa" |
305 assume d: "\<exists>xa\<in>set a. x \<in> set xa" |
373 show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d]) |
306 show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d]) |
374 next |
307 next |
375 assume e: "\<exists>xa\<in>set b. x \<in> set xa" |
308 assume e: "\<exists>xa\<in>set b. x \<in> set xa" |
376 have a': "list_all2 op \<approx> ba a" by (rule list_all2_symp[OF list_eq_equivp, OF a]) |
309 have a': "list_all2 op \<approx> ba a" by (rule list_all2_symp[OF list_eq_equivp, OF a]) |
377 have b': "bb \<approx> ba" by (rule equivp_symp[OF list_eq_equivp, OF b]) |
310 have b': "bb \<approx> ba" by (rule equivp_symp[OF list_eq_equivp, OF b]) |
378 have c': "list_all2 op \<approx> b bb" by (rule list_all2_symp[OF list_eq_equivp, OF c]) |
311 have c': "list_all2 op \<approx> b bb" by (rule list_all2_symp[OF list_eq_equivp, OF c]) |
379 show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e]) |
312 show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e]) |
380 qed |
313 qed |
381 qed |
314 qed |
382 then show "concat a \<approx> concat b" by simp |
315 then show "concat a \<approx> concat b" by auto |
383 qed |
316 qed |
384 |
|
385 |
317 |
386 |
318 |
387 lemma concat_rsp_unfolded: |
319 lemma concat_rsp_unfolded: |
388 "\<lbrakk>list_all2 op \<approx> a ba; ba \<approx> bb; list_all2 op \<approx> bb b\<rbrakk> \<Longrightarrow> concat a \<approx> concat b" |
320 "\<lbrakk>list_all2 op \<approx> a ba; ba \<approx> bb; list_all2 op \<approx> bb b\<rbrakk> \<Longrightarrow> concat a \<approx> concat b" |
389 proof - |
321 proof - |
485 f_subset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) |
415 f_subset_eq :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subseteq>|" 50) |
486 where |
416 where |
487 "xs |\<subseteq>| ys \<equiv> xs \<le> ys" |
417 "xs |\<subseteq>| ys \<equiv> xs \<le> ys" |
488 |
418 |
489 definition |
419 definition |
490 less_fset: |
420 less_fset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" |
491 "(xs :: 'a fset) < ys \<equiv> xs \<le> ys \<and> xs \<noteq> ys" |
421 where |
|
422 "xs < ys \<equiv> xs \<le> ys \<and> xs \<noteq> (ys::'a fset)" |
492 |
423 |
493 abbreviation |
424 abbreviation |
494 f_subset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) |
425 fsubset :: "'a fset \<Rightarrow> 'a fset \<Rightarrow> bool" (infix "|\<subset>|" 50) |
495 where |
426 where |
496 "xs |\<subset>| ys \<equiv> xs < ys" |
427 "xs |\<subset>| ys \<equiv> xs < ys" |
497 |
428 |
498 quotient_definition |
429 quotient_definition |
499 "sup \<Colon> ('a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset)" |
430 "sup :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
500 is |
431 is |
501 "(op @) \<Colon> ('a list \<Rightarrow> 'a list \<Rightarrow> 'a list)" |
432 "append :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
502 |
433 |
503 abbreviation |
434 abbreviation |
504 funion (infixl "|\<union>|" 65) |
435 funion (infixl "|\<union>|" 65) |
505 where |
436 where |
506 "xs |\<union>| ys \<equiv> sup (xs :: 'a fset) ys" |
437 "xs |\<union>| ys \<equiv> sup xs (ys::'a fset)" |
507 |
438 |
508 quotient_definition |
439 quotient_definition |
509 "inf \<Colon> ('a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset)" |
440 "inf :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
510 is |
441 is |
511 "finter_raw \<Colon> ('a list \<Rightarrow> 'a list \<Rightarrow> 'a list)" |
442 "finter_raw :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
512 |
443 |
513 abbreviation |
444 abbreviation |
514 finter (infixl "|\<inter>|" 65) |
445 finter (infixl "|\<inter>|" 65) |
515 where |
446 where |
516 "xs |\<inter>| ys \<equiv> inf (xs :: 'a fset) ys" |
447 "xs |\<inter>| ys \<equiv> inf xs (ys::'a fset)" |
517 |
448 |
518 quotient_definition |
449 quotient_definition |
519 "minus \<Colon> ('a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset)" |
450 "minus :: 'a fset \<Rightarrow> 'a fset \<Rightarrow> 'a fset" |
520 is |
451 is |
521 "fminus_raw \<Colon> ('a list \<Rightarrow> 'a list \<Rightarrow> 'a list)" |
452 "fminus_raw :: 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" |
522 |
453 |
523 instance |
454 instance |
524 proof |
455 proof |
525 fix x y z :: "'a fset" |
456 fix x y z :: "'a fset" |
526 show "(x |\<subset>| y) = (x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x)" |
457 show "(x |\<subset>| y) = (x |\<subseteq>| y \<and> \<not> y |\<subseteq>| x)" |
527 unfolding less_fset by (lifting sub_list_not_eq) |
458 unfolding less_fset_def by (lifting sub_list_not_eq) |
528 show "x |\<subseteq>| x" by (lifting sub_list_refl) |
459 show "x |\<subseteq>| x" by (lifting sub_list_refl) |
529 show "{||} |\<subseteq>| x" by (lifting sub_list_empty) |
460 show "{||} |\<subseteq>| x" by (lifting sub_list_empty) |
530 show "x |\<subseteq>| x |\<union>| y" by (lifting sub_list_append_left) |
461 show "x |\<subseteq>| x |\<union>| y" by (lifting sub_list_append_left) |
531 show "y |\<subseteq>| x |\<union>| y" by (lifting sub_list_append_right) |
462 show "y |\<subseteq>| x |\<union>| y" by (lifting sub_list_append_right) |
532 show "x |\<inter>| y |\<subseteq>| x" by (lifting sub_list_inter_left) |
463 show "x |\<inter>| y |\<subseteq>| x" by (lifting sub_list_inter_left) |
721 shows "memb x xs \<Longrightarrow> memb x (y # xs)" |
649 shows "memb x xs \<Longrightarrow> memb x (y # xs)" |
722 by (simp add: memb_def) |
650 by (simp add: memb_def) |
723 |
651 |
724 lemma singleton_list_eq: |
652 lemma singleton_list_eq: |
725 shows "[x] \<approx> [y] \<longleftrightarrow> x = y" |
653 shows "[x] \<approx> [y] \<longleftrightarrow> x = y" |
726 by (simp add:) auto |
654 by (simp) |
727 |
655 |
728 lemma sub_list_cons: |
656 lemma sub_list_cons: |
729 "sub_list (x # xs) ys = (memb x ys \<and> sub_list xs ys)" |
657 "sub_list (x # xs) ys = (memb x ys \<and> sub_list xs ys)" |
730 by (auto simp add: memb_def sub_list_def) |
658 by (auto simp add: memb_def sub_list_def) |
731 |
659 |
732 lemma fminus_raw_red: "fminus_raw (x # xs) ys = (if memb x ys then fminus_raw xs ys else x # (fminus_raw xs ys))" |
660 lemma fminus_raw_red: |
733 by (induct ys arbitrary: xs x) |
661 "fminus_raw (x # xs) ys = (if x \<in> set ys then fminus_raw xs ys else x # (fminus_raw xs ys))" |
734 (simp_all add: not_memb_nil memb_delete_raw memb_cons_iff) |
662 by (induct ys arbitrary: xs x) (simp_all) |
735 |
663 |
736 text {* Cardinality of finite sets *} |
664 text {* Cardinality of finite sets *} |
737 |
665 |
|
666 (* used in memb_card_not_0 *) |
738 lemma fcard_raw_0: |
667 lemma fcard_raw_0: |
739 shows "fcard_raw xs = 0 \<longleftrightarrow> xs \<approx> []" |
668 shows "fcard_raw xs = 0 \<longleftrightarrow> xs \<approx> []" |
740 by (induct xs) (auto simp add: memb_def) |
669 unfolding fcard_raw_def |
741 |
670 by (induct xs) (auto) |
742 lemma fcard_raw_not_memb: |
671 |
743 shows "\<not> memb x xs \<longleftrightarrow> fcard_raw (x # xs) = Suc (fcard_raw xs)" |
672 (* used in list_eq2_equiv *) |
744 by auto |
|
745 |
|
746 lemma fcard_raw_suc: |
|
747 assumes a: "fcard_raw xs = Suc n" |
|
748 shows "\<exists>x ys. \<not> (memb x ys) \<and> xs \<approx> (x # ys) \<and> fcard_raw ys = n" |
|
749 using a |
|
750 by (induct xs) (auto simp add: memb_def split_ifs) |
|
751 |
|
752 lemma singleton_fcard_1: |
|
753 shows "set xs = {x} \<Longrightarrow> fcard_raw xs = 1" |
|
754 by (induct xs) (auto simp add: memb_def subset_insert) |
|
755 |
|
756 lemma fcard_raw_1: |
|
757 shows "fcard_raw xs = 1 \<longleftrightarrow> (\<exists>x. xs \<approx> [x])" |
|
758 apply (auto dest!: fcard_raw_suc) |
|
759 apply (simp add: fcard_raw_0) |
|
760 apply (rule_tac x="x" in exI) |
|
761 apply simp |
|
762 apply (subgoal_tac "set xs = {x}") |
|
763 apply (drule singleton_fcard_1) |
|
764 apply auto |
|
765 done |
|
766 |
|
767 lemma fcard_raw_suc_memb: |
|
768 assumes a: "fcard_raw A = Suc n" |
|
769 shows "\<exists>a. memb a A" |
|
770 using a |
|
771 by (induct A) (auto simp add: memb_def) |
|
772 |
|
773 lemma memb_card_not_0: |
673 lemma memb_card_not_0: |
774 assumes a: "memb a A" |
674 assumes a: "memb a A" |
775 shows "\<not>(fcard_raw A = 0)" |
675 shows "\<not>(fcard_raw A = 0)" |
776 proof - |
676 proof - |
777 have "\<not>(\<forall>x. \<not> memb x A)" using a by auto |
677 have "\<not>(\<forall>x. \<not> memb x A)" using a by auto |
778 then have "\<not>A \<approx> []" using none_memb_nil[of A] by simp |
678 then have "\<not>A \<approx> []" using none_memb_nil[of A] by simp |
779 then show ?thesis using fcard_raw_0[of A] by simp |
679 then show ?thesis using fcard_raw_0[of A] by simp |
780 qed |
680 qed |
781 |
681 |
782 text {* fmap *} |
682 |
|
683 |
|
684 section {* fmap *} |
|
685 |
|
686 (* there is another fmap section below *) |
783 |
687 |
784 lemma map_append: |
688 lemma map_append: |
785 "map f (xs @ ys) \<approx> (map f xs) @ (map f ys)" |
689 "map f (xs @ ys) \<approx> (map f xs) @ (map f ys)" |
786 by simp |
690 by simp |
787 |
691 |
828 qed |
732 qed |
829 qed |
733 qed |
830 then show thesis using a c by blast |
734 then show thesis using a c by blast |
831 qed |
735 qed |
832 |
736 |
|
737 |
833 section {* deletion *} |
738 section {* deletion *} |
834 |
739 |
835 lemma memb_delete_raw_ident: |
740 lemma memb_removeAll_ident: |
836 shows "\<not> memb x (delete_raw xs x)" |
741 shows "\<not> memb x (removeAll x xs)" |
837 by (induct xs) (auto simp add: memb_def) |
742 by (induct xs) (auto simp add: memb_def) |
838 |
743 |
839 lemma fset_raw_delete_raw_cases: |
744 lemma fset_raw_removeAll_cases: |
840 "xs = [] \<or> (\<exists>x. memb x xs \<and> xs \<approx> x # delete_raw xs x)" |
745 "xs = [] \<or> (\<exists>x. memb x xs \<and> xs \<approx> x # removeAll x xs)" |
841 by (induct xs) (auto simp add: memb_def) |
746 by (induct xs) (auto simp add: memb_def) |
842 |
747 |
843 lemma fdelete_raw_filter: |
748 lemma fremoveAll_filter: |
844 "delete_raw xs y = [x \<leftarrow> xs. x \<noteq> y]" |
749 "removeAll y xs = [x \<leftarrow> xs. x \<noteq> y]" |
845 by (induct xs) simp_all |
750 by (induct xs) simp_all |
846 |
751 |
847 lemma fcard_raw_delete: |
752 lemma fcard_raw_delete: |
848 "fcard_raw (delete_raw xs y) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)" |
753 "fcard_raw (removeAll y xs) = (if memb y xs then fcard_raw xs - 1 else fcard_raw xs)" |
849 by (simp add: fdelete_raw_filter fcard_raw_delete_one) |
754 by (auto simp add: fcard_raw_def memb_def) |
850 |
755 |
851 lemma finter_raw_empty: |
756 lemma finter_raw_empty: |
852 "finter_raw l [] = []" |
757 "finter_raw l [] = []" |
853 by (induct l) (simp_all add: not_memb_nil) |
758 by (induct l) (simp_all add: not_memb_nil) |
854 |
|
855 lemma set_cong: |
|
856 shows "(x \<approx> y) = (set x = set y)" |
|
857 by auto |
|
858 |
759 |
859 lemma inj_map_eq_iff: |
760 lemma inj_map_eq_iff: |
860 "inj f \<Longrightarrow> (map f l \<approx> map f m) = (l \<approx> m)" |
761 "inj f \<Longrightarrow> (map f l \<approx> map f m) = (l \<approx> m)" |
861 by (simp add: set_eq_iff[symmetric] inj_image_eq_iff) |
762 by (simp add: set_eq_iff[symmetric] inj_image_eq_iff) |
862 |
763 |
876 lemma list_eq2_refl: |
777 lemma list_eq2_refl: |
877 shows "list_eq2 xs xs" |
778 shows "list_eq2 xs xs" |
878 by (induct xs) (auto intro: list_eq2.intros) |
779 by (induct xs) (auto intro: list_eq2.intros) |
879 |
780 |
880 lemma cons_delete_list_eq2: |
781 lemma cons_delete_list_eq2: |
881 shows "list_eq2 (a # (delete_raw A a)) (if memb a A then A else a # A)" |
782 shows "list_eq2 (a # (removeAll a A)) (if memb a A then A else a # A)" |
882 apply (induct A) |
783 apply (induct A) |
883 apply (simp add: memb_def list_eq2_refl) |
784 apply (simp add: memb_def list_eq2_refl) |
884 apply (case_tac "memb a (aa # A)") |
785 apply (case_tac "memb a (aa # A)") |
885 apply (simp_all only: memb_cons_iff) |
786 apply (simp_all only: memb_cons_iff) |
886 apply (case_tac [!] "a = aa") |
787 apply (case_tac [!] "a = aa") |
887 apply (simp_all) |
788 apply (simp_all) |
888 apply (case_tac "memb a A") |
789 apply (case_tac "memb a A") |
889 apply (auto simp add: memb_def)[2] |
790 apply (auto simp add: memb_def)[2] |
890 apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6)) |
791 apply (metis list_eq2.intros(3) list_eq2.intros(4) list_eq2.intros(5) list_eq2.intros(6)) |
891 apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6)) |
792 apply (metis list_eq2.intros(1) list_eq2.intros(5) list_eq2.intros(6)) |
892 apply (auto simp add: list_eq2_refl not_memb_delete_raw_ident) |
793 apply (auto simp add: list_eq2_refl not_memb_removeAll_ident) |
893 done |
794 done |
894 |
795 |
895 lemma memb_delete_list_eq2: |
796 lemma memb_delete_list_eq2: |
896 assumes a: "memb e r" |
797 assumes a: "memb e r" |
897 shows "list_eq2 (e # delete_raw r e) r" |
798 shows "list_eq2 (e # removeAll e r) r" |
898 using a cons_delete_list_eq2[of e r] |
799 using a cons_delete_list_eq2[of e r] |
899 by simp |
800 by simp |
900 |
801 |
901 lemma delete_raw_rsp: |
802 lemma removeAll_rsp: |
902 "xs \<approx> ys \<Longrightarrow> delete_raw xs x \<approx> delete_raw ys x" |
803 "xs \<approx> ys \<Longrightarrow> removeAll x xs \<approx> removeAll x ys" |
903 by (simp add: memb_def[symmetric] memb_delete_raw) |
804 by (simp add: memb_def[symmetric]) |
904 |
805 |
905 lemma list_eq2_equiv: |
806 lemma list_eq2_equiv: |
906 "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)" |
807 "(l \<approx> r) \<longleftrightarrow> (list_eq2 l r)" |
907 proof |
808 proof |
908 show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto |
809 show "list_eq2 l r \<Longrightarrow> l \<approx> r" by (induct rule: list_eq2.induct) auto |
921 then show ?case using z list_eq2_refl by simp |
822 then show ?case using z list_eq2_refl by simp |
922 next |
823 next |
923 case (Suc m) |
824 case (Suc m) |
924 have b: "l \<approx> r" by fact |
825 have b: "l \<approx> r" by fact |
925 have d: "fcard_raw l = Suc m" by fact |
826 have d: "fcard_raw l = Suc m" by fact |
926 then have "\<exists>a. memb a l" by (rule fcard_raw_suc_memb) |
827 then have "\<exists>a. memb a l" |
|
828 apply(simp add: fcard_raw_def memb_def) |
|
829 apply(drule card_eq_SucD) |
|
830 apply(blast) |
|
831 done |
927 then obtain a where e: "memb a l" by auto |
832 then obtain a where e: "memb a l" by auto |
928 then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b by auto |
833 then have e': "memb a r" using list_eq.simps[simplified memb_def[symmetric], of l r] b |
929 have f: "fcard_raw (delete_raw l a) = m" using fcard_raw_delete[of l a] e d by simp |
834 unfolding memb_def by auto |
930 have g: "delete_raw l a \<approx> delete_raw r a" using delete_raw_rsp[OF b] by simp |
835 have f: "fcard_raw (removeAll a l) = m" using fcard_raw_delete[of a l] e d by simp |
931 have "list_eq2 (delete_raw l a) (delete_raw r a)" by (rule Suc.hyps[OF f g]) |
836 have g: "removeAll a l \<approx> removeAll a r" using removeAll_rsp[OF b] by simp |
932 then have h: "list_eq2 (a # delete_raw l a) (a # delete_raw r a)" by (rule list_eq2.intros(5)) |
837 have "list_eq2 (removeAll a l) (removeAll a r)" by (rule Suc.hyps[OF f g]) |
933 have i: "list_eq2 l (a # delete_raw l a)" |
838 then have h: "list_eq2 (a # removeAll a l) (a # removeAll a r)" by (rule list_eq2.intros(5)) |
|
839 have i: "list_eq2 l (a # removeAll a l)" |
934 by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]]) |
840 by (rule list_eq2.intros(3)[OF memb_delete_list_eq2[OF e]]) |
935 have "list_eq2 l (a # delete_raw r a)" by (rule list_eq2.intros(6)[OF i h]) |
841 have "list_eq2 l (a # removeAll a r)" by (rule list_eq2.intros(6)[OF i h]) |
936 then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp |
842 then show ?case using list_eq2.intros(6)[OF _ memb_delete_list_eq2[OF e']] by simp |
937 qed |
843 qed |
938 } |
844 } |
939 then show "l \<approx> r \<Longrightarrow> list_eq2 l r" by blast |
845 then show "l \<approx> r \<Longrightarrow> list_eq2 l r" by blast |
940 qed |
846 qed |
941 |
847 |
942 text {* Set *} |
|
943 |
|
944 lemma sub_list_set: "sub_list xs ys = (set xs \<subseteq> set ys)" |
|
945 by (metis rev_append set_append set_cong set_rev sub_list_append sub_list_append_left sub_list_def sub_list_not_eq subset_Un_eq) |
|
946 |
|
947 lemma sub_list_neq_set: "(sub_list xs ys \<and> \<not> list_eq xs ys) = (set xs \<subset> set ys)" |
|
948 by (auto simp add: sub_list_set) |
|
949 |
|
950 lemma fcard_raw_set: "fcard_raw xs = card (set xs)" |
|
951 by (induct xs) (auto simp add: insert_absorb memb_def card_insert_disjoint) |
|
952 |
|
953 lemma memb_set: "memb x xs = (x \<in> set xs)" |
|
954 by (simp only: memb_def) |
|
955 |
|
956 lemma filter_set: "set (filter P xs) = P \<inter> (set xs)" |
|
957 by (induct xs, simp) |
|
958 (metis Int_insert_right_if0 Int_insert_right_if1 List.set.simps(2) filter.simps(2) mem_def) |
|
959 |
|
960 lemma delete_raw_set: "set (delete_raw xs x) = set xs - {x}" |
|
961 by (induct xs) auto |
|
962 |
|
963 lemma inter_raw_set: "set (finter_raw xs ys) = set xs \<inter> set ys" |
|
964 by (induct xs) (simp_all add: memb_def) |
|
965 |
|
966 lemma fminus_raw_set: "set (fminus_raw xs ys) = set xs - set ys" |
|
967 by (induct ys arbitrary: xs) |
|
968 (simp_all add: delete_raw_set, blast) |
|
969 |
|
970 text {* Raw theorems of ffilter *} |
|
971 |
|
972 lemma sub_list_filter: "sub_list (filter P xs) (filter Q xs) = (\<forall> x. memb x xs \<longrightarrow> P x \<longrightarrow> Q x)" |
|
973 unfolding sub_list_def memb_def by auto |
|
974 |
|
975 lemma list_eq_filter: "list_eq (filter P xs) (filter Q xs) = (\<forall>x. memb x xs \<longrightarrow> P x = Q x)" |
|
976 unfolding memb_def by auto |
|
977 |
|
978 text {* Lifted theorems *} |
848 text {* Lifted theorems *} |
979 |
849 |
980 lemma not_fin_fnil: "x |\<notin>| {||}" |
850 lemma not_fin_fnil: "x |\<notin>| {||}" |
981 by (lifting not_memb_nil) |
851 by (lifting not_memb_nil) |
982 |
852 |
1008 |
878 |
1009 lemma fsingleton_eq[simp]: |
879 lemma fsingleton_eq[simp]: |
1010 shows "{|x|} = {|y|} \<longleftrightarrow> x = y" |
880 shows "{|x|} = {|y|} \<longleftrightarrow> x = y" |
1011 by (lifting singleton_list_eq) |
881 by (lifting singleton_list_eq) |
1012 |
882 |
1013 text {* fset_to_set *} |
883 |
1014 |
884 text {* fset *} |
1015 lemma fset_to_set_simps[simp]: |
885 |
1016 "fset_to_set {||} = ({} :: 'a set)" |
886 lemma fset_simps[simp]: |
1017 "fset_to_set (finsert (h :: 'a) t) = insert h (fset_to_set t)" |
887 "fset {||} = ({} :: 'a set)" |
|
888 "fset (finsert (h :: 'a) t) = insert h (fset t)" |
1018 by (lifting set.simps) |
889 by (lifting set.simps) |
1019 |
890 |
1020 lemma in_fset_to_set: |
891 lemma in_fset: |
1021 "x \<in> fset_to_set S \<equiv> x |\<in>| S" |
892 "x \<in> fset S \<equiv> x |\<in>| S" |
1022 by (lifting memb_def[symmetric]) |
893 by (lifting memb_def[symmetric]) |
1023 |
894 |
1024 lemma none_fin_fempty: |
895 lemma none_fin_fempty: |
1025 "(\<forall>x. x |\<notin>| S) = (S = {||})" |
896 "(\<forall>x. x |\<notin>| S) = (S = {||})" |
1026 by (lifting none_memb_nil) |
897 by (lifting none_memb_nil) |
1027 |
898 |
1028 lemma fset_cong: |
899 lemma fset_cong: |
1029 "(S = T) = (fset_to_set S = fset_to_set T)" |
900 "(S = T) = (fset S = fset T)" |
1030 by (lifting set_cong) |
901 by (lifting list_eq.simps) |
1031 |
902 |
1032 text {* fcard *} |
903 |
1033 |
904 section {* fcard *} |
1034 lemma fcard_fempty [simp]: |
|
1035 shows "fcard {||} = 0" |
|
1036 by (lifting fcard_raw_nil) |
|
1037 |
905 |
1038 lemma fcard_finsert_if [simp]: |
906 lemma fcard_finsert_if [simp]: |
1039 shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))" |
907 shows "fcard (finsert x S) = (if x |\<in>| S then fcard S else Suc (fcard S))" |
1040 by (lifting fcard_raw_cons) |
908 by (descending) (auto simp add: fcard_raw_def memb_def insert_absorb) |
1041 |
909 |
1042 lemma fcard_0: "(fcard S = 0) = (S = {||})" |
910 lemma fcard_0[simp]: |
1043 by (lifting fcard_raw_0) |
911 shows "fcard S = 0 \<longleftrightarrow> S = {||}" |
|
912 by (descending) (simp add: fcard_raw_def) |
|
913 |
|
914 lemma fcard_fempty[simp]: |
|
915 shows "fcard {||} = 0" |
|
916 by (simp add: fcard_0) |
1044 |
917 |
1045 lemma fcard_1: |
918 lemma fcard_1: |
1046 shows "(fcard S = 1) = (\<exists>x. S = {|x|})" |
919 shows "fcard S = 1 \<longleftrightarrow> (\<exists>x. S = {|x|})" |
1047 by (lifting fcard_raw_1) |
920 by (descending) (auto simp add: fcard_raw_def card_Suc_eq) |
1048 |
921 |
1049 lemma fcard_gt_0: |
922 lemma fcard_gt_0: |
1050 shows "x \<in> fset_to_set S \<Longrightarrow> 0 < fcard S" |
923 shows "x \<in> fset S \<Longrightarrow> 0 < fcard S" |
1051 by (lifting fcard_raw_gt_0) |
924 by (descending) (auto simp add: fcard_raw_def card_gt_0_iff) |
1052 |
925 |
1053 lemma fcard_not_fin: |
926 lemma fcard_not_fin: |
1054 shows "(x |\<notin>| S) = (fcard (finsert x S) = Suc (fcard S))" |
927 assumes a: "x |\<notin>| S" |
1055 by (lifting fcard_raw_not_memb) |
928 shows "fcard (finsert x S) = Suc (fcard S)" |
1056 |
929 using a |
1057 lemma fcard_suc: "fcard S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = finsert x T \<and> fcard T = n" |
930 by (descending) (simp add: memb_def fcard_raw_def) |
1058 by (lifting fcard_raw_suc) |
931 |
|
932 lemma fcard_suc: |
|
933 shows "fcard S = Suc n \<Longrightarrow> \<exists>x T. x |\<notin>| T \<and> S = finsert x T \<and> fcard T = n" |
|
934 apply(descending) |
|
935 apply(simp add: fcard_raw_def memb_def) |
|
936 apply(drule card_eq_SucD) |
|
937 apply(auto) |
|
938 apply(rule_tac x="b" in exI) |
|
939 apply(rule_tac x="removeAll b S" in exI) |
|
940 apply(auto) |
|
941 done |
1059 |
942 |
1060 lemma fcard_delete: |
943 lemma fcard_delete: |
1061 "fcard (fdelete S y) = (if y |\<in>| S then fcard S - 1 else fcard S)" |
944 "fcard (fdelete y S) = (if y |\<in>| S then fcard S - 1 else fcard S)" |
1062 by (lifting fcard_raw_delete) |
945 by (lifting fcard_raw_delete) |
1063 |
946 |
1064 lemma fcard_suc_memb: "fcard A = Suc n \<Longrightarrow> \<exists>a. a |\<in>| A" |
947 lemma fcard_suc_memb: |
1065 by (lifting fcard_raw_suc_memb) |
948 shows "fcard A = Suc n \<Longrightarrow> \<exists>a. a |\<in>| A" |
1066 |
949 apply(descending) |
1067 lemma fin_fcard_not_0: "a |\<in>| A \<Longrightarrow> fcard A \<noteq> 0" |
950 apply(simp add: fcard_raw_def memb_def) |
1068 by (lifting memb_card_not_0) |
951 apply(drule card_eq_SucD) |
1069 |
952 apply(auto) |
1070 text {* funion *} |
953 done |
|
954 |
|
955 lemma fin_fcard_not_0: |
|
956 shows "a |\<in>| A \<Longrightarrow> fcard A \<noteq> 0" |
|
957 by (descending) (auto simp add: fcard_raw_def memb_def) |
|
958 |
|
959 |
|
960 section {* funion *} |
1071 |
961 |
1072 lemmas [simp] = |
962 lemmas [simp] = |
1073 sup_bot_left[where 'a="'a fset", standard] |
963 sup_bot_left[where 'a="'a fset", standard] |
1074 sup_bot_right[where 'a="'a fset", standard] |
964 sup_bot_right[where 'a="'a fset", standard] |
1075 |
965 |
1076 lemma funion_finsert[simp]: |
966 lemma funion_finsert[simp]: |
1077 shows "finsert x S |\<union>| T = finsert x (S |\<union>| T)" |
967 shows "finsert x S |\<union>| T = finsert x (S |\<union>| T)" |
1078 by (lifting append.simps(2)) |
968 by (lifting append.simps(2)) |
1079 |
969 |
1080 lemma singleton_union_left: |
970 lemma singleton_union_left: |
1081 "{|a|} |\<union>| S = finsert a S" |
971 shows "{|a|} |\<union>| S = finsert a S" |
1082 by simp |
972 by simp |
1083 |
973 |
1084 lemma singleton_union_right: |
974 lemma singleton_union_right: |
1085 "S |\<union>| {|a|} = finsert a S" |
975 shows "S |\<union>| {|a|} = finsert a S" |
1086 by (subst sup.commute) simp |
976 by (subst sup.commute) simp |
1087 |
977 |
1088 section {* Induction and Cases rules for finite sets *} |
978 |
|
979 section {* Induction and Cases rules for fsets *} |
1089 |
980 |
1090 lemma fset_strong_cases: |
981 lemma fset_strong_cases: |
1091 obtains "xs = {||}" |
982 obtains "xs = {||}" |
1092 | x ys where "x |\<notin>| ys" and "xs = finsert x ys" |
983 | x ys where "x |\<notin>| ys" and "xs = finsert x ys" |
1093 by (lifting fset_raw_strong_cases) |
984 by (lifting fset_raw_strong_cases) |
1161 |
1053 |
1162 lemma fin_funion: |
1054 lemma fin_funion: |
1163 "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T" |
1055 "x |\<in>| S |\<union>| T \<longleftrightarrow> x |\<in>| S \<or> x |\<in>| T" |
1164 by (lifting memb_append) |
1056 by (lifting memb_append) |
1165 |
1057 |
1166 text {* to_set *} |
1058 |
1167 |
1059 section {* fset *} |
1168 lemma fin_set: "(x |\<in>| xs) = (x \<in> fset_to_set xs)" |
1060 |
1169 by (lifting memb_set) |
1061 lemma fin_set: |
1170 |
1062 shows "x |\<in>| xs \<longleftrightarrow> x \<in> fset xs" |
1171 lemma fnotin_set: "(x |\<notin>| xs) = (x \<notin> fset_to_set xs)" |
1063 by (lifting memb_def) |
|
1064 |
|
1065 lemma fnotin_set: |
|
1066 shows "x |\<notin>| xs \<longleftrightarrow> x \<notin> fset xs" |
1172 by (simp add: fin_set) |
1067 by (simp add: fin_set) |
1173 |
1068 |
1174 lemma fcard_set: "fcard xs = card (fset_to_set xs)" |
1069 lemma fcard_set: |
1175 by (lifting fcard_raw_set) |
1070 shows "fcard xs = card (fset xs)" |
1176 |
1071 by (lifting fcard_raw_def) |
1177 lemma fsubseteq_set: "(xs |\<subseteq>| ys) = (fset_to_set xs \<subseteq> fset_to_set ys)" |
1072 |
1178 by (lifting sub_list_set) |
1073 lemma fsubseteq_set: |
1179 |
1074 shows "xs |\<subseteq>| ys \<longleftrightarrow> fset xs \<subseteq> fset ys" |
1180 lemma fsubset_set: "(xs |\<subset>| ys) = (fset_to_set xs \<subset> fset_to_set ys)" |
1075 by (lifting sub_list_def) |
1181 unfolding less_fset by (lifting sub_list_neq_set) |
1076 |
1182 |
1077 lemma fsubset_set: |
1183 lemma ffilter_set: "fset_to_set (ffilter P xs) = P \<inter> fset_to_set xs" |
1078 shows "xs |\<subset>| ys \<longleftrightarrow> fset xs \<subset> fset ys" |
1184 by (lifting filter_set) |
1079 unfolding less_fset_def |
1185 |
1080 by (descending) (auto simp add: sub_list_def) |
1186 lemma fdelete_set: "fset_to_set (fdelete xs x) = fset_to_set xs - {x}" |
1081 |
1187 by (lifting delete_raw_set) |
1082 lemma ffilter_set [simp]: |
1188 |
1083 shows "fset (ffilter P xs) = P \<inter> fset xs" |
1189 lemma inter_set: "fset_to_set (xs |\<inter>| ys) = fset_to_set xs \<inter> fset_to_set ys" |
1084 by (descending) (auto simp add: mem_def) |
1190 by (lifting inter_raw_set) |
1085 |
1191 |
1086 lemma fdelete_set [simp]: |
1192 lemma union_set: "fset_to_set (xs |\<union>| ys) = fset_to_set xs \<union> fset_to_set ys" |
1087 shows "fset (fdelete x xs) = fset xs - {x}" |
|
1088 by (lifting set_removeAll) |
|
1089 |
|
1090 lemma finter_set [simp]: |
|
1091 shows "fset (xs |\<inter>| ys) = fset xs \<inter> fset ys" |
|
1092 by (lifting set_finter_raw) |
|
1093 |
|
1094 lemma funion_set [simp]: |
|
1095 shows "fset (xs |\<union>| ys) = fset xs \<union> fset ys" |
1193 by (lifting set_append) |
1096 by (lifting set_append) |
1194 |
1097 |
1195 lemma fminus_set: "fset_to_set (xs - ys) = fset_to_set xs - fset_to_set ys" |
1098 lemma fminus_set [simp]: |
1196 by (lifting fminus_raw_set) |
1099 shows "fset (xs - ys) = fset xs - fset ys" |
1197 |
1100 by (lifting set_fminus_raw) |
1198 lemmas fset_to_set_trans = |
1101 |
1199 fin_set fnotin_set fcard_set fsubseteq_set fsubset_set |
1102 |
1200 inter_set union_set ffilter_set fset_to_set_simps |
1103 |
1201 fset_cong fdelete_set fmap_set_image fminus_set |
1104 section {* ffold *} |
1202 |
1105 |
1203 |
1106 lemma ffold_nil: |
1204 text {* ffold *} |
1107 shows "ffold f z {||} = z" |
1205 |
|
1206 lemma ffold_nil: "ffold f z {||} = z" |
|
1207 by (lifting ffold_raw.simps(1)[where 'a="'b" and 'b="'a"]) |
1108 by (lifting ffold_raw.simps(1)[where 'a="'b" and 'b="'a"]) |
1208 |
1109 |
1209 lemma ffold_finsert: "ffold f z (finsert a A) = |
1110 lemma ffold_finsert: "ffold f z (finsert a A) = |
1210 (if rsp_fold f then if a |\<in>| A then ffold f z A else f a (ffold f z A) else z)" |
1111 (if rsp_fold f then if a |\<in>| A then ffold f z A else f a (ffold f z A) else z)" |
1211 by (lifting ffold_raw.simps(2)[where 'a="'b" and 'b="'a"]) |
1112 by (descending) (simp add: memb_def) |
1212 |
1113 |
1213 lemma fin_commute_ffold: |
1114 lemma fin_commute_ffold: |
1214 "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> ffold f z b = f h (ffold f z (fdelete b h))" |
1115 "\<lbrakk>rsp_fold f; h |\<in>| b\<rbrakk> \<Longrightarrow> ffold f z b = f h (ffold f z (fdelete h b))" |
1215 by (lifting memb_commute_ffold_raw) |
1116 by (descending) (simp add: memb_def memb_commute_ffold_raw) |
1216 |
1117 |
1217 text {* fdelete *} |
1118 |
|
1119 |
|
1120 section {* fdelete *} |
1218 |
1121 |
1219 lemma fin_fdelete: |
1122 lemma fin_fdelete: |
1220 shows "x |\<in>| fdelete S y \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y" |
1123 shows "x |\<in>| fdelete y S \<longleftrightarrow> x |\<in>| S \<and> x \<noteq> y" |
1221 by (lifting memb_delete_raw) |
1124 by (descending) (simp add: memb_def) |
1222 |
1125 |
1223 lemma fin_fdelete_ident: |
1126 lemma fin_fdelete_ident: |
1224 shows "x |\<notin>| fdelete S x" |
1127 shows "x |\<notin>| fdelete x S" |
1225 by (lifting memb_delete_raw_ident) |
1128 by (lifting memb_removeAll_ident) |
1226 |
1129 |
1227 lemma not_memb_fdelete_ident: |
1130 lemma not_memb_fdelete_ident: |
1228 shows "x |\<notin>| S \<Longrightarrow> fdelete S x = S" |
1131 shows "x |\<notin>| S \<Longrightarrow> fdelete x S = S" |
1229 by (lifting not_memb_delete_raw_ident) |
1132 by (lifting not_memb_removeAll_ident) |
1230 |
1133 |
1231 lemma fset_fdelete_cases: |
1134 lemma fset_fdelete_cases: |
1232 shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = finsert x (fdelete S x))" |
1135 shows "S = {||} \<or> (\<exists>x. x |\<in>| S \<and> S = finsert x (fdelete x S))" |
1233 by (lifting fset_raw_delete_raw_cases) |
1136 by (lifting fset_raw_removeAll_cases) |
1234 |
1137 |
1235 text {* inter *} |
1138 |
|
1139 section {* finter *} |
1236 |
1140 |
1237 lemma finter_empty_l: "({||} |\<inter>| S) = {||}" |
1141 lemma finter_empty_l: "({||} |\<inter>| S) = {||}" |
1238 by (lifting finter_raw.simps(1)) |
1142 by (lifting finter_raw.simps(1)) |
1239 |
1143 |
1240 lemma finter_empty_r: "(S |\<inter>| {||}) = {||}" |
1144 lemma finter_empty_r: "(S |\<inter>| {||}) = {||}" |
1241 by (lifting finter_raw_empty) |
1145 by (lifting finter_raw_empty) |
1242 |
1146 |
1243 lemma finter_finsert: |
1147 lemma finter_finsert: |
1244 "finsert x S |\<inter>| T = (if x |\<in>| T then finsert x (S |\<inter>| T) else S |\<inter>| T)" |
1148 shows "finsert x S |\<inter>| T = (if x |\<in>| T then finsert x (S |\<inter>| T) else S |\<inter>| T)" |
1245 by (lifting finter_raw.simps(2)) |
1149 by (descending) (simp add: memb_def) |
1246 |
1150 |
1247 lemma fin_finter: |
1151 lemma fin_finter: |
1248 "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T" |
1152 shows "x |\<in>| (S |\<inter>| T) \<longleftrightarrow> x |\<in>| S \<and> x |\<in>| T" |
1249 by (lifting memb_finter_raw) |
1153 by (descending) (simp add: memb_def) |
1250 |
1154 |
1251 lemma fsubset_finsert: |
1155 lemma fsubset_finsert: |
1252 "(finsert x xs |\<subseteq>| ys) = (x |\<in>| ys \<and> xs |\<subseteq>| ys)" |
1156 shows "(finsert x xs |\<subseteq>| ys) = (x |\<in>| ys \<and> xs |\<subseteq>| ys)" |
1253 by (lifting sub_list_cons) |
1157 by (lifting sub_list_cons) |
1254 |
1158 |
1255 lemma "xs |\<subseteq>| ys \<equiv> \<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys" |
1159 lemma |
1256 by (lifting sub_list_def[simplified memb_def[symmetric]]) |
1160 shows "xs |\<subseteq>| ys \<equiv> \<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys" |
1257 |
1161 by (descending) (auto simp add: sub_list_def memb_def) |
1258 lemma fsubset_fin: "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)" |
1162 |
1259 by (rule meta_eq_to_obj_eq) |
1163 lemma fsubset_fin: |
1260 (lifting sub_list_def[simplified memb_def[symmetric]]) |
1164 shows "xs |\<subseteq>| ys = (\<forall>x. x |\<in>| xs \<longrightarrow> x |\<in>| ys)" |
1261 |
1165 by (descending) (auto simp add: sub_list_def memb_def) |
1262 lemma fminus_fin: "(x |\<in>| xs - ys) = (x |\<in>| xs \<and> x |\<notin>| ys)" |
1166 |
1263 by (lifting fminus_raw_memb) |
1167 lemma fminus_fin: |
1264 |
1168 shows "(x |\<in>| xs - ys) = (x |\<in>| xs \<and> x |\<notin>| ys)" |
1265 lemma fminus_red: "finsert x xs - ys = (if x |\<in>| ys then xs - ys else finsert x (xs - ys))" |
1169 by (descending) (simp add: memb_def) |
1266 by (lifting fminus_raw_red) |
1170 |
1267 |
1171 lemma fminus_red: |
1268 lemma fminus_red_fin[simp]: "x |\<in>| ys \<Longrightarrow> finsert x xs - ys = xs - ys" |
1172 shows "finsert x xs - ys = (if x |\<in>| ys then xs - ys else finsert x (xs - ys))" |
|
1173 by (descending) (auto simp add: memb_def) |
|
1174 |
|
1175 lemma fminus_red_fin[simp]: |
|
1176 shows "x |\<in>| ys \<Longrightarrow> finsert x xs - ys = xs - ys" |
1269 by (simp add: fminus_red) |
1177 by (simp add: fminus_red) |
1270 |
1178 |
1271 lemma fminus_red_fnotin[simp]: "x |\<notin>| ys \<Longrightarrow> finsert x xs - ys = finsert x (xs - ys)" |
1179 lemma fminus_red_fnotin[simp]: |
|
1180 shows "x |\<notin>| ys \<Longrightarrow> finsert x xs - ys = finsert x (xs - ys)" |
1272 by (simp add: fminus_red) |
1181 by (simp add: fminus_red) |
1273 |
1182 |
1274 lemma fset_eq_iff: |
1183 lemma fset_eq_iff: |
1275 "(S = T) = (\<forall>x. (x |\<in>| S) = (x |\<in>| T))" |
1184 "(S = T) = (\<forall>x. (x |\<in>| S) = (x |\<in>| T))" |
1276 by (lifting list_eq.simps[simplified memb_def[symmetric]]) |
1185 by (descending) (auto simp add: memb_def) |
1277 |
1186 |
1278 (* We cannot write it as "assumes .. shows" since Isabelle changes |
1187 (* We cannot write it as "assumes .. shows" since Isabelle changes |
1279 the quantifiers to schematic variables and reintroduces them in |
1188 the quantifiers to schematic variables and reintroduces them in |
1280 a different order *) |
1189 a different order *) |
1281 lemma fset_eq_cases: |
1190 lemma fset_eq_cases: |
1298 and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3" |
1207 and "\<And>xs1 xs2 xs3. \<lbrakk>xs1 = xs2; P xs1 xs2; xs2 = xs3; P xs2 xs3\<rbrakk> \<Longrightarrow> P xs1 xs3" |
1299 shows "P x1 x2" |
1208 shows "P x1 x2" |
1300 using assms |
1209 using assms |
1301 by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]]) |
1210 by (lifting list_eq2.induct[simplified list_eq2_equiv[symmetric]]) |
1302 |
1211 |
1303 text {* concat *} |
1212 |
|
1213 section {* fconcat *} |
1304 |
1214 |
1305 lemma fconcat_empty: |
1215 lemma fconcat_empty: |
1306 shows "fconcat {||} = {||}" |
1216 shows "fconcat {||} = {||}" |
1307 by (lifting concat.simps(1)) |
1217 by (lifting concat.simps(1)) |
1308 |
1218 |
1309 lemma fconcat_insert: |
1219 lemma fconcat_insert: |
1310 shows "fconcat (finsert x S) = x |\<union>| fconcat S" |
1220 shows "fconcat (finsert x S) = x |\<union>| fconcat S" |
1311 by (lifting concat.simps(2)) |
1221 by (lifting concat.simps(2)) |
1312 |
1222 |
1313 text {* ffilter *} |
1223 |
1314 |
1224 section {* ffilter *} |
1315 lemma subseteq_filter: "ffilter P xs <= ffilter Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)" |
1225 |
1316 by (lifting sub_list_filter) |
1226 lemma subseteq_filter: |
1317 |
1227 shows "ffilter P xs <= ffilter Q xs = (\<forall> x. x |\<in>| xs \<longrightarrow> P x \<longrightarrow> Q x)" |
1318 lemma eq_ffilter: "(ffilter P xs = ffilter Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)" |
1228 by (descending) (auto simp add: memb_def sub_list_def) |
1319 by (lifting list_eq_filter) |
1229 |
1320 |
1230 lemma eq_ffilter: |
|
1231 shows "(ffilter P xs = ffilter Q xs) = (\<forall>x. x |\<in>| xs \<longrightarrow> P x = Q x)" |
|
1232 by (descending) (auto simp add: memb_def) |
|
1233 |
1321 lemma subset_ffilter: |
1234 lemma subset_ffilter: |
1322 "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> ffilter P xs < ffilter Q xs" |
1235 "(\<And>x. x |\<in>| xs \<Longrightarrow> P x \<Longrightarrow> Q x) \<Longrightarrow> (x |\<in>| xs & \<not> P x & Q x) \<Longrightarrow> ffilter P xs < ffilter Q xs" |
1323 unfolding less_fset by (auto simp add: subseteq_filter eq_ffilter) |
1236 unfolding less_fset_def by (auto simp add: subseteq_filter eq_ffilter) |
|
1237 |
1324 |
1238 |
1325 section {* lemmas transferred from Finite_Set theory *} |
1239 section {* lemmas transferred from Finite_Set theory *} |
1326 |
1240 |
1327 text {* finiteness for finite sets holds *} |
1241 text {* finiteness for finite sets holds *} |
1328 lemma finite_fset: "finite (fset_to_set S)" |
1242 lemma finite_fset [simp]: |
|
1243 shows "finite (fset S)" |
1329 by (induct S) auto |
1244 by (induct S) auto |
1330 |
1245 |
1331 lemma fset_choice: "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)" |
1246 lemma fset_choice: |
1332 unfolding fset_to_set_trans |
1247 shows "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>y. P x y) \<Longrightarrow> \<exists>f. \<forall>x. x |\<in>| A \<longrightarrow> P x (f x)" |
1333 by (rule finite_set_choice[simplified Ball_def, OF finite_fset]) |
1248 apply(descending) |
1334 |
1249 apply(simp add: memb_def) |
1335 lemma fsubseteq_fnil: "xs |\<subseteq>| {||} = (xs = {||})" |
1250 apply(rule finite_set_choice[simplified Ball_def]) |
1336 unfolding fset_to_set_trans |
1251 apply(simp_all) |
1337 by (rule subset_empty) |
1252 done |
1338 |
1253 |
1339 lemma not_fsubset_fnil: "\<not> xs |\<subset>| {||}" |
1254 lemma fsubseteq_fempty: |
1340 unfolding fset_to_set_trans |
1255 shows "xs |\<subseteq>| {||} = (xs = {||})" |
1341 by (rule not_psubset_empty) |
1256 by (metis finter_empty_r le_iff_inf) |
1342 |
1257 |
1343 lemma fcard_mono: "xs |\<subseteq>| ys \<Longrightarrow> fcard xs \<le> fcard ys" |
1258 lemma not_fsubset_fnil: |
1344 unfolding fset_to_set_trans |
1259 shows "\<not> xs |\<subset>| {||}" |
|
1260 by (metis fset_simps(1) fsubset_set not_psubset_empty) |
|
1261 |
|
1262 lemma fcard_mono: |
|
1263 shows "xs |\<subseteq>| ys \<Longrightarrow> fcard xs \<le> fcard ys" |
|
1264 unfolding fcard_set fsubseteq_set |
1345 by (rule card_mono[OF finite_fset]) |
1265 by (rule card_mono[OF finite_fset]) |
1346 |
1266 |
1347 lemma fcard_fseteq: "xs |\<subseteq>| ys \<Longrightarrow> fcard ys \<le> fcard xs \<Longrightarrow> xs = ys" |
1267 lemma fcard_fseteq: |
1348 unfolding fset_to_set_trans |
1268 shows "xs |\<subseteq>| ys \<Longrightarrow> fcard ys \<le> fcard xs \<Longrightarrow> xs = ys" |
1349 by (rule card_seteq[OF finite_fset]) |
1269 unfolding fcard_set fsubseteq_set |
1350 |
1270 by (simp add: card_seteq[OF finite_fset] fset_cong) |
1351 lemma psubset_fcard_mono: "xs |\<subset>| ys \<Longrightarrow> fcard xs < fcard ys" |
1271 |
1352 unfolding fset_to_set_trans |
1272 lemma psubset_fcard_mono: |
|
1273 shows "xs |\<subset>| ys \<Longrightarrow> fcard xs < fcard ys" |
|
1274 unfolding fcard_set fsubset_set |
1353 by (rule psubset_card_mono[OF finite_fset]) |
1275 by (rule psubset_card_mono[OF finite_fset]) |
1354 |
1276 |
1355 lemma fcard_funion_finter: "fcard xs + fcard ys = fcard (xs |\<union>| ys) + fcard (xs |\<inter>| ys)" |
1277 lemma fcard_funion_finter: |
1356 unfolding fset_to_set_trans |
1278 shows "fcard xs + fcard ys = fcard (xs |\<union>| ys) + fcard (xs |\<inter>| ys)" |
|
1279 unfolding fcard_set funion_set finter_set |
1357 by (rule card_Un_Int[OF finite_fset finite_fset]) |
1280 by (rule card_Un_Int[OF finite_fset finite_fset]) |
1358 |
1281 |
1359 lemma fcard_funion_disjoint: "xs |\<inter>| ys = {||} \<Longrightarrow> fcard (xs |\<union>| ys) = fcard xs + fcard ys" |
1282 lemma fcard_funion_disjoint: |
1360 unfolding fset_to_set_trans |
1283 shows "xs |\<inter>| ys = {||} \<Longrightarrow> fcard (xs |\<union>| ys) = fcard xs + fcard ys" |
1361 by (rule card_Un_disjoint[OF finite_fset finite_fset]) |
1284 unfolding fcard_set funion_set |
1362 |
1285 apply (rule card_Un_disjoint[OF finite_fset finite_fset]) |
1363 lemma fcard_delete1_less: "x |\<in>| xs \<Longrightarrow> fcard (fdelete xs x) < fcard xs" |
1286 by (metis finter_set fset_simps(1)) |
1364 unfolding fset_to_set_trans |
1287 |
|
1288 lemma fcard_delete1_less: |
|
1289 shows "x |\<in>| xs \<Longrightarrow> fcard (fdelete x xs) < fcard xs" |
|
1290 unfolding fcard_set fin_set fdelete_set |
1365 by (rule card_Diff1_less[OF finite_fset]) |
1291 by (rule card_Diff1_less[OF finite_fset]) |
1366 |
1292 |
1367 lemma fcard_delete2_less: "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> fcard (fdelete (fdelete xs x) y) < fcard xs" |
1293 lemma fcard_delete2_less: |
1368 unfolding fset_to_set_trans |
1294 shows "x |\<in>| xs \<Longrightarrow> y |\<in>| xs \<Longrightarrow> fcard (fdelete y (fdelete x xs)) < fcard xs" |
|
1295 unfolding fcard_set fdelete_set fin_set |
1369 by (rule card_Diff2_less[OF finite_fset]) |
1296 by (rule card_Diff2_less[OF finite_fset]) |
1370 |
1297 |
1371 lemma fcard_delete1_le: "fcard (fdelete xs x) <= fcard xs" |
1298 lemma fcard_delete1_le: |
1372 unfolding fset_to_set_trans |
1299 shows "fcard (fdelete x xs) <= fcard xs" |
|
1300 unfolding fdelete_set fcard_set |
1373 by (rule card_Diff1_le[OF finite_fset]) |
1301 by (rule card_Diff1_le[OF finite_fset]) |
1374 |
1302 |
1375 lemma fcard_psubset: "ys |\<subseteq>| xs \<Longrightarrow> fcard ys < fcard xs \<Longrightarrow> ys |\<subset>| xs" |
1303 lemma fcard_psubset: |
1376 unfolding fset_to_set_trans |
1304 shows "ys |\<subseteq>| xs \<Longrightarrow> fcard ys < fcard xs \<Longrightarrow> ys |\<subset>| xs" |
|
1305 unfolding fcard_set fsubseteq_set fsubset_set |
1377 by (rule card_psubset[OF finite_fset]) |
1306 by (rule card_psubset[OF finite_fset]) |
1378 |
1307 |
1379 lemma fcard_fmap_le: "fcard (fmap f xs) \<le> fcard xs" |
1308 lemma fcard_fmap_le: |
1380 unfolding fset_to_set_trans |
1309 shows "fcard (fmap f xs) \<le> fcard xs" |
|
1310 unfolding fcard_set fmap_set_image |
1381 by (rule card_image_le[OF finite_fset]) |
1311 by (rule card_image_le[OF finite_fset]) |
1382 |
1312 |
1383 lemma fin_fminus_fnotin: "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S" |
1313 lemma fin_fminus_fnotin: |
1384 unfolding fset_to_set_trans |
1314 shows "x |\<in>| F - S \<Longrightarrow> x |\<notin>| S" |
|
1315 unfolding fin_set fminus_set |
1385 by blast |
1316 by blast |
1386 |
1317 |
1387 lemma fin_fnotin_fminus: "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S" |
1318 lemma fin_fnotin_fminus: |
1388 unfolding fset_to_set_trans |
1319 shows "x |\<in>| S \<Longrightarrow> x |\<notin>| F - S" |
|
1320 unfolding fin_set fminus_set |
1389 by blast |
1321 by blast |
1390 |
1322 |
1391 lemma fin_mdef: "x |\<in>| F = ((x |\<notin>| (F - {|x|})) & (F = finsert x (F - {|x|})))" |
1323 lemma fin_mdef: |
1392 unfolding fset_to_set_trans |
1324 shows "x |\<in>| F = ((x |\<notin>| (F - {|x|})) & (F = finsert x (F - {|x|})))" |
|
1325 unfolding fin_set fset_simps fset_cong fminus_set |
1393 by blast |
1326 by blast |
1394 |
1327 |
1395 lemma fcard_fminus_finsert[simp]: |
1328 lemma fcard_fminus_finsert[simp]: |
1396 assumes "a |\<in>| A" and "a |\<notin>| B" |
1329 assumes "a |\<in>| A" and "a |\<notin>| B" |
1397 shows "fcard(A - finsert a B) = fcard(A - B) - 1" |
1330 shows "fcard (A - finsert a B) = fcard (A - B) - 1" |
1398 using assms unfolding fset_to_set_trans |
1331 using assms |
1399 by (rule card_Diff_insert[OF finite_fset]) |
1332 unfolding fin_set fcard_set fminus_set |
|
1333 by (simp add: card_Diff_insert[OF finite_fset]) |
1400 |
1334 |
1401 lemma fcard_fminus_fsubset: |
1335 lemma fcard_fminus_fsubset: |
1402 assumes "B |\<subseteq>| A" |
1336 assumes "B |\<subseteq>| A" |
1403 shows "fcard (A - B) = fcard A - fcard B" |
1337 shows "fcard (A - B) = fcard A - fcard B" |
1404 using assms unfolding fset_to_set_trans |
1338 using assms |
|
1339 unfolding fsubseteq_set fcard_set fminus_set |
1405 by (rule card_Diff_subset[OF finite_fset]) |
1340 by (rule card_Diff_subset[OF finite_fset]) |
1406 |
1341 |
1407 lemma fcard_fminus_subset_finter: |
1342 lemma fcard_fminus_subset_finter: |
1408 "fcard (A - B) = fcard A - fcard (A |\<inter>| B)" |
1343 shows "fcard (A - B) = fcard A - fcard (A |\<inter>| B)" |
1409 unfolding fset_to_set_trans |
1344 using assms |
1410 by (rule card_Diff_subset_Int) (fold inter_set, rule finite_fset) |
1345 unfolding finter_set fcard_set fminus_set |
1411 |
1346 by (rule card_Diff_subset_Int) (simp) |
1412 lemma ball_reg_right_unfolded: "(\<forall>x. R x \<longrightarrow> P x \<longrightarrow> Q x) \<longrightarrow> (All P \<longrightarrow> Ball R Q)" |
1347 |
1413 apply rule |
|
1414 apply (rule ball_reg_right) |
|
1415 apply auto |
|
1416 done |
|
1417 |
1348 |
1418 lemma list_all2_refl: |
1349 lemma list_all2_refl: |
1419 assumes q: "equivp R" |
1350 assumes q: "equivp R" |
1420 shows "(list_all2 R) r r" |
1351 shows "(list_all2 R) r r" |
1421 by (rule list_all2_refl) (metis equivp_def q) |
1352 by (rule list_all2_refl) (metis equivp_def q) |