| author | wenzelm | 
| Fri, 14 Dec 2012 16:33:22 +0100 | |
| changeset 50530 | 6266e44b3396 | 
| parent 50144 | 885deccc264e | 
| child 51371 | 197ad6b8f763 | 
| permissions | -rw-r--r-- | 
| 49509 
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
 blanchet parents: 
49507diff
changeset | 1 | (* Title: HOL/BNF/More_BNFs.thy | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 2 | Author: Dmitriy Traytel, TU Muenchen | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 3 | Author: Andrei Popescu, TU Muenchen | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 4 | Author: Andreas Lochbihler, Karlsruhe Institute of Technology | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 5 | Author: Jasmin Blanchette, TU Muenchen | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 6 | Copyright 2012 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 7 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 8 | Registration of various types as bounded natural functors. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 9 | *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 10 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 11 | header {* Registration of Various Types as Bounded Natural Functors *}
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 12 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 13 | theory More_BNFs | 
| 49310 | 14 | imports | 
| 15 | BNF_LFP | |
| 16 | BNF_GFP | |
| 17 | "~~/src/HOL/Quotient_Examples/FSet" | |
| 18 | "~~/src/HOL/Library/Multiset" | |
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50027diff
changeset | 19 | Countable_Type | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 20 | begin | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 21 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 22 | lemma option_rec_conv_option_case: "option_rec = option_case" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 23 | by (simp add: fun_eq_iff split: option.split) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 24 | |
| 49507 | 25 | definition option_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a option \<Rightarrow> 'b option \<Rightarrow> bool" where
 | 
| 26 | "option_rel R x_opt y_opt = | |
| 49463 | 27 | (case (x_opt, y_opt) of | 
| 28 | (None, None) \<Rightarrow> True | |
| 29 | | (Some x, Some y) \<Rightarrow> R x y | |
| 30 | | _ \<Rightarrow> False)" | |
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 31 | |
| 49507 | 32 | bnf_def Option.map [Option.set] "\<lambda>_::'a option. natLeq" ["None"] option_rel | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 33 | proof - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 34 | show "Option.map id = id" by (simp add: fun_eq_iff Option.map_def split: option.split) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 35 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 36 | fix f g | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 37 | show "Option.map (g \<circ> f) = Option.map g \<circ> Option.map f" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 38 | by (auto simp add: fun_eq_iff Option.map_def split: option.split) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 39 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 40 | fix f g x | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 41 | assume "\<And>z. z \<in> Option.set x \<Longrightarrow> f z = g z" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 42 | thus "Option.map f x = Option.map g x" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 43 | by (simp cong: Option.map_cong) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 44 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 45 | fix f | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 46 | show "Option.set \<circ> Option.map f = op ` f \<circ> Option.set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 47 | by fastforce | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 48 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 49 | show "card_order natLeq" by (rule natLeq_card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 50 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 51 | show "cinfinite natLeq" by (rule natLeq_cinfinite) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 52 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 53 | fix x | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 54 | show "|Option.set x| \<le>o natLeq" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 55 | by (cases x) (simp_all add: ordLess_imp_ordLeq finite_iff_ordLess_natLeq[symmetric]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 56 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 57 | fix A | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 58 |   have unfold: "{x. Option.set x \<subseteq> A} = Some ` A \<union> {None}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 59 | by (auto simp add: option_rec_conv_option_case Option.set_def split: option.split_asm) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 60 |   show "|{x. Option.set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 61 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 62 | apply (rule card_of_ordIso_subst[OF unfold]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 63 | apply (rule ordLeq_transitive) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 64 | apply (rule Un_csum) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 65 | apply (rule ordLeq_transitive) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 66 | apply (rule csum_mono) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 67 | apply (rule card_of_image) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 68 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 69 | apply (rule single_cone) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 70 | apply (rule cone_ordLeq_ctwo) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 71 | apply (rule ordLeq_cexp1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 72 | apply (simp_all add: natLeq_cinfinite natLeq_Card_order cinfinite_not_czero Card_order_csum) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 73 | done | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 74 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 75 | fix A B1 B2 f1 f2 p1 p2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 76 | assume wpull: "wpull A B1 B2 f1 f2 p1 p2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 77 |   show "wpull {x. Option.set x \<subseteq> A} {x. Option.set x \<subseteq> B1} {x. Option.set x \<subseteq> B2}
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 78 | (Option.map f1) (Option.map f2) (Option.map p1) (Option.map p2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 79 | (is "wpull ?A ?B1 ?B2 ?f1 ?f2 ?p1 ?p2") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 80 | unfolding wpull_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 81 | proof (intro strip, elim conjE) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 82 | fix b1 b2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 83 | assume "b1 \<in> ?B1" "b2 \<in> ?B2" "?f1 b1 = ?f2 b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 84 | thus "\<exists>a \<in> ?A. ?p1 a = b1 \<and> ?p2 a = b2" using wpull | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 85 | unfolding wpull_def by (cases b2) (auto 4 5) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 86 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 87 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 88 | fix z | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 89 | assume "z \<in> Option.set None" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 90 | thus False by simp | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 91 | next | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 92 | fix R | 
| 49507 | 93 |   show "{p. option_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
 | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 94 |         (Gr {x. Option.set x \<subseteq> R} (Option.map fst))\<inverse> O Gr {x. Option.set x \<subseteq> R} (Option.map snd)"
 | 
| 49507 | 95 | unfolding option_rel_def Gr_def relcomp_unfold converse_unfold | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 96 | by (auto simp: trans[OF eq_commute option_map_is_None] trans[OF eq_commute option_map_eq_Some] | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 97 | split: option.splits) blast | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 98 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 99 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 100 | lemma card_of_list_in: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 101 |   "|{xs. set xs \<subseteq> A}| \<le>o |Pfunc (UNIV :: nat set) A|" (is "|?LHS| \<le>o |?RHS|")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 102 | proof - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 103 | let ?f = "%xs. %i. if i < length xs \<and> set xs \<subseteq> A then Some (nth xs i) else None" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 104 | have "inj_on ?f ?LHS" unfolding inj_on_def fun_eq_iff | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 105 | proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 106 | fix xs :: "'a list" and ys :: "'a list" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 107 | assume su: "set xs \<subseteq> A" "set ys \<subseteq> A" and eq: "\<forall>i. ?f xs i = ?f ys i" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 108 | hence *: "length xs = length ys" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 109 | by (metis linorder_cases option.simps(2) order_less_irrefl) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 110 | thus "xs = ys" by (rule nth_equalityI) (metis * eq su option.inject) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 111 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 112 | moreover have "?f ` ?LHS \<subseteq> ?RHS" unfolding Pfunc_def by fastforce | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 113 | ultimately show ?thesis using card_of_ordLeq by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 114 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 115 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 116 | lemma list_in_empty: "A = {} \<Longrightarrow> {x. set x \<subseteq> A} = {[]}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 117 | by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 118 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 119 | lemma card_of_Func: "|Func A B| =o |B| ^c |A|" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 120 | unfolding cexp_def Field_card_of by (rule card_of_refl) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 121 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 122 | lemma not_emp_czero_notIn_ordIso_Card_order: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 123 | "A \<noteq> {} \<Longrightarrow> ( |A|, czero) \<notin> ordIso \<and> Card_order |A|"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 124 | apply (rule conjI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 125 | apply (metis Field_card_of czeroE) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 126 | by (rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 127 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 128 | lemma list_in_bd: "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 129 | proof - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 130 | fix A :: "'a set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 131 |   show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 132 |   proof (cases "A = {}")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 133 | case False thus ?thesis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 134 | apply - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 135 | apply (rule ordLeq_transitive) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 136 | apply (rule card_of_list_in) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 137 | apply (rule ordLeq_transitive) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 138 | apply (erule card_of_Pfunc_Pow_Func) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 139 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 140 | apply (rule Times_cprod) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 141 | apply (rule cprod_cinfinite_bound) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 142 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 143 | apply (rule Pow_cexp_ctwo) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 144 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 145 | apply (rule cexp_cong2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 146 | apply (rule card_of_nat) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 147 | apply (rule Card_order_ctwo) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 148 | apply (rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 149 | apply (rule natLeq_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 150 | apply (rule disjI1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 151 | apply (rule ctwo_Cnotzero) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 152 | apply (rule cexp_mono1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 153 | apply (rule ordLeq_csum2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 154 | apply (rule Card_order_ctwo) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 155 | apply (rule disjI1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 156 | apply (rule ctwo_Cnotzero) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 157 | apply (rule natLeq_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 158 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 159 | apply (rule card_of_Func) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 160 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 161 | apply (rule cexp_cong2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 162 | apply (rule card_of_nat) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 163 | apply (rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 164 | apply (rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 165 | apply (rule natLeq_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 166 | apply (rule disjI1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 167 | apply (erule not_emp_czero_notIn_ordIso_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 168 | apply (rule cexp_mono1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 169 | apply (rule ordLeq_csum1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 170 | apply (rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 171 | apply (rule disjI1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 172 | apply (erule not_emp_czero_notIn_ordIso_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 173 | apply (rule natLeq_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 174 | apply (rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 175 | apply (rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 176 | apply (rule Cinfinite_cexp) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 177 | apply (rule ordLeq_csum2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 178 | apply (rule Card_order_ctwo) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 179 | apply (rule conjI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 180 | apply (rule natLeq_cinfinite) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 181 | by (rule natLeq_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 182 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 183 | case True thus ?thesis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 184 | apply - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 185 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 186 | apply (rule card_of_ordIso_subst) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 187 | apply (erule list_in_empty) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 188 | apply (rule ordIso_ordLeq_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 189 | apply (rule single_cone) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 190 | apply (rule cone_ordLeq_cexp) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 191 | apply (rule ordLeq_transitive) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 192 | apply (rule cone_ordLeq_ctwo) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 193 | apply (rule ordLeq_csum2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 194 | by (rule Card_order_ctwo) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 195 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 196 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 197 | |
| 49434 
433dc7e028c8
separated registration of BNFs from bnf_def (BNFs are now stored only for bnf_def and (co)data commands)
 traytel parents: 
49316diff
changeset | 198 | bnf_def map [set] "\<lambda>_::'a list. natLeq" ["[]"] | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 199 | proof - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 200 | show "map id = id" by (rule List.map.id) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 201 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 202 | fix f g | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 203 | show "map (g o f) = map g o map f" by (rule List.map.comp[symmetric]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 204 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 205 | fix x f g | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 206 | assume "\<And>z. z \<in> set x \<Longrightarrow> f z = g z" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 207 | thus "map f x = map g x" by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 208 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 209 | fix f | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 210 | show "set o map f = image f o set" by (rule ext, unfold o_apply, rule set_map) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 211 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 212 | show "card_order natLeq" by (rule natLeq_card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 213 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 214 | show "cinfinite natLeq" by (rule natLeq_cinfinite) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 215 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 216 | fix x | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 217 | show "|set x| \<le>o natLeq" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 218 | apply (rule ordLess_imp_ordLeq) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 219 | apply (rule finite_ordLess_infinite[OF _ natLeq_Well_order]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 220 | unfolding Field_natLeq Field_card_of by (auto simp: card_of_well_order_on) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 221 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 222 | fix A :: "'a set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 223 |   show "|{x. set x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd)
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 224 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 225 | fix A B1 B2 f1 f2 p1 p2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 226 | assume "wpull A B1 B2 f1 f2 p1 p2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 227 | hence pull: "\<And>b1 b2. b1 \<in> B1 \<and> b2 \<in> B2 \<and> f1 b1 = f2 b2 \<Longrightarrow> \<exists>a \<in> A. p1 a = b1 \<and> p2 a = b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 228 | unfolding wpull_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 229 |   show "wpull {x. set x \<subseteq> A} {x. set x \<subseteq> B1} {x. set x \<subseteq> B2} (map f1) (map f2) (map p1) (map p2)"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 230 | (is "wpull ?A ?B1 ?B2 _ _ _ _") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 231 | proof (unfold wpull_def) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 232 |     { fix as bs assume *: "as \<in> ?B1" "bs \<in> ?B2" "map f1 as = map f2 bs"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 233 | hence "length as = length bs" by (metis length_map) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 234 | hence "\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs" using * | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 235 | proof (induct as bs rule: list_induct2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 236 | case (Cons a as b bs) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 237 | hence "a \<in> B1" "b \<in> B2" "f1 a = f2 b" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 238 | with pull obtain z where "z \<in> A" "p1 z = a" "p2 z = b" by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 239 | moreover | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 240 | from Cons obtain zs where "zs \<in> ?A" "map p1 zs = as" "map p2 zs = bs" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 241 | ultimately have "z # zs \<in> ?A" "map p1 (z # zs) = a # as \<and> map p2 (z # zs) = b # bs" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 242 | thus ?case by (rule_tac x = "z # zs" in bexI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 243 | qed simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 244 | } | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 245 | thus "\<forall>as bs. as \<in> ?B1 \<and> bs \<in> ?B2 \<and> map f1 as = map f2 bs \<longrightarrow> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 246 | (\<exists>zs \<in> ?A. map p1 zs = as \<and> map p2 zs = bs)" by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 247 | qed | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 248 | qed simp+ | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 249 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 250 | (* Finite sets *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 251 | abbreviation afset where "afset \<equiv> abs_fset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 252 | abbreviation rfset where "rfset \<equiv> rep_fset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 253 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 254 | lemma fset_fset_member: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 255 | "fset A = {a. a |\<in>| A}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 256 | unfolding fset_def fset_member_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 257 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 258 | lemma afset_rfset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 259 | "afset (rfset x) = x" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 260 | by (rule Quotient_fset[unfolded Quotient_def, THEN conjunct1, rule_format]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 261 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 262 | lemma afset_rfset_id: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 263 | "afset o rfset = id" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 264 | unfolding comp_def afset_rfset id_def .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 265 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 266 | lemma rfset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 267 | "rfset A = rfset B \<longleftrightarrow> A = B" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 268 | by (metis afset_rfset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 269 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 270 | lemma afset_set: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 271 | "afset as = afset bs \<longleftrightarrow> set as = set bs" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 272 | using Quotient_fset unfolding Quotient_def list_eq_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 273 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 274 | lemma surj_afset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 275 | "\<exists> as. A = afset as" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 276 | by (metis afset_rfset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 277 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 278 | lemma fset_def2: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 279 | "fset = set o rfset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 280 | unfolding fset_def map_fun_def[abs_def] by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 281 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 282 | lemma fset_def2_raw: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 283 | "fset A = set (rfset A)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 284 | unfolding fset_def2 by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 285 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 286 | lemma fset_comp_afset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 287 | "fset o afset = set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 288 | unfolding fset_def2 comp_def apply(rule ext) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 289 | unfolding afset_set[symmetric] afset_rfset .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 290 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 291 | lemma fset_afset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 292 | "fset (afset as) = set as" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 293 | unfolding fset_comp_afset[symmetric] by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 294 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 295 | lemma set_rfset_afset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 296 | "set (rfset (afset as)) = set as" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 297 | unfolding afset_set[symmetric] afset_rfset .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 298 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 299 | lemma map_fset_comp_afset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 300 | "(map_fset f) o afset = afset o (map f)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 301 | unfolding map_fset_def map_fun_def[abs_def] comp_def apply(rule ext) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 302 | unfolding afset_set set_map set_rfset_afset id_apply .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 303 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 304 | lemma map_fset_afset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 305 | "(map_fset f) (afset as) = afset (map f as)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 306 | using map_fset_comp_afset unfolding comp_def fun_eq_iff by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 307 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 308 | lemma fset_map_fset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 309 | "fset (map_fset f A) = (image f) (fset A)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 310 | apply(subst afset_rfset[symmetric, of A]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 311 | unfolding map_fset_afset fset_afset set_map | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 312 | unfolding fset_def2_raw .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 313 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 314 | lemma map_fset_def2: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 315 | "map_fset f = afset o (map f) o rfset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 316 | unfolding map_fset_def map_fun_def[abs_def] by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 317 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 318 | lemma map_fset_def2_raw: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 319 | "map_fset f A = afset (map f (rfset A))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 320 | unfolding map_fset_def2 by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 321 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 322 | lemma finite_ex_fset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 323 | assumes "finite A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 324 | shows "\<exists> B. fset B = A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 325 | by (metis assms finite_list fset_afset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 326 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 327 | lemma wpull_image: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 328 | assumes "wpull A B1 B2 f1 f2 p1 p2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 329 | shows "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 330 | unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 331 | fix Y1 Y2 assume Y1: "Y1 \<subseteq> B1" and Y2: "Y2 \<subseteq> B2" and EQ: "f1 ` Y1 = f2 ` Y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 332 |   def X \<equiv> "{a \<in> A. p1 a \<in> Y1 \<and> p2 a \<in> Y2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 333 | show "\<exists>X\<subseteq>A. p1 ` X = Y1 \<and> p2 ` X = Y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 334 | proof (rule exI[of _ X], intro conjI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 335 | show "p1 ` X = Y1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 336 | proof | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 337 | show "Y1 \<subseteq> p1 ` X" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 338 | proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 339 | fix y1 assume y1: "y1 \<in> Y1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 340 | then obtain y2 where y2: "y2 \<in> Y2" and eq: "f1 y1 = f2 y2" using EQ by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 341 | then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 342 | using assms y1 Y1 Y2 unfolding wpull_def by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 343 | thus "y1 \<in> p1 ` X" unfolding X_def using y1 y2 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 344 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 345 | qed(unfold X_def, auto) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 346 | show "p2 ` X = Y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 347 | proof | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 348 | show "Y2 \<subseteq> p2 ` X" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 349 | proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 350 | fix y2 assume y2: "y2 \<in> Y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 351 | then obtain y1 where y1: "y1 \<in> Y1" and eq: "f1 y1 = f2 y2" using EQ by force | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 352 | then obtain x where "x \<in> A" and "p1 x = y1" and "p2 x = y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 353 | using assms y2 Y1 Y2 unfolding wpull_def by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 354 | thus "y2 \<in> p2 ` X" unfolding X_def using y1 y2 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 355 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 356 | qed(unfold X_def, auto) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 357 | qed(unfold X_def, auto) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 358 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 359 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 360 | lemma fset_to_fset: "finite A \<Longrightarrow> fset (the_inv fset A) = A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 361 | by (rule f_the_inv_into_f) (auto simp: inj_on_def fset_cong dest!: finite_ex_fset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 362 | |
| 49507 | 363 | definition fset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a fset \<Rightarrow> 'b fset \<Rightarrow> bool" where
 | 
| 364 | "fset_rel R a b \<longleftrightarrow> | |
| 49463 | 365 | (\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> | 
| 366 | (\<forall>t \<in> fset b. \<exists>u \<in> fset a. R u t)" | |
| 367 | ||
| 49507 | 368 | lemma fset_rel_aux: | 
| 49463 | 369 | "(\<forall>t \<in> fset a. \<exists>u \<in> fset b. R t u) \<and> (\<forall>u \<in> fset b. \<exists>t \<in> fset a. R t u) \<longleftrightarrow> | 
| 370 |  (a, b) \<in> (Gr {a. fset a \<subseteq> {(a, b). R a b}} (map_fset fst))\<inverse> O
 | |
| 371 |           Gr {a. fset a \<subseteq> {(a, b). R a b}} (map_fset snd)" (is "?L = ?R")
 | |
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 372 | proof | 
| 49463 | 373 | assume ?L | 
| 374 | def R' \<equiv> "the_inv fset (Collect (split R) \<inter> (fset a \<times> fset b))" (is "the_inv fset ?L'") | |
| 375 | have "finite ?L'" by (intro finite_Int[OF disjI2] finite_cartesian_product) auto | |
| 376 | hence *: "fset R' = ?L'" unfolding R'_def by (intro fset_to_fset) | |
| 377 | show ?R unfolding Gr_def relcomp_unfold converse_unfold | |
| 378 | proof (intro CollectI prod_caseI exI conjI) | |
| 379 | from * show "(R', a) = (R', map_fset fst R')" using conjunct1[OF `?L`] | |
| 380 | by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits) | |
| 381 | from * show "(R', b) = (R', map_fset snd R')" using conjunct2[OF `?L`] | |
| 382 | by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits) | |
| 383 | qed (auto simp add: *) | |
| 384 | next | |
| 385 | assume ?R thus ?L unfolding Gr_def relcomp_unfold converse_unfold | |
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 386 | apply (simp add: subset_eq Ball_def) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 387 | apply (rule conjI) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 388 | apply (clarsimp, metis snd_conv) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 389 | by (clarsimp, metis fst_conv) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 390 | qed | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 391 | |
| 49507 | 392 | bnf_def map_fset [fset] "\<lambda>_::'a fset. natLeq" ["{||}"] fset_rel
 | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 393 | proof - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 394 | show "map_fset id = id" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 395 | unfolding map_fset_def2 map_id o_id afset_rfset_id .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 396 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 397 | fix f g | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 398 | show "map_fset (g o f) = map_fset g o map_fset f" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 399 | unfolding map_fset_def2 map.comp[symmetric] comp_def apply(rule ext) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 400 | unfolding afset_set set_map fset_def2_raw[symmetric] image_image[symmetric] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 401 | unfolding map_fset_afset[symmetric] map_fset_image afset_rfset | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 402 | by (rule refl) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 403 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 404 | fix x f g | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 405 | assume "\<And>z. z \<in> fset x \<Longrightarrow> f z = g z" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 406 | hence "map f (rfset x) = map g (rfset x)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 407 | apply(intro map_cong) unfolding fset_def2_raw by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 408 | thus "map_fset f x = map_fset g x" unfolding map_fset_def2_raw | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 409 | by (rule arg_cong) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 410 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 411 | fix f | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 412 | show "fset o map_fset f = image f o fset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 413 | unfolding comp_def fset_map_fset .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 414 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 415 | show "card_order natLeq" by (rule natLeq_card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 416 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 417 | show "cinfinite natLeq" by (rule natLeq_cinfinite) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 418 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 419 | fix x | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 420 | show "|fset x| \<le>o natLeq" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 421 | unfolding fset_def2_raw | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 422 | apply (rule ordLess_imp_ordLeq) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 423 | apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 424 | by (rule finite_set) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 425 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 426 | fix A :: "'a set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 427 |   have "|{x. fset x \<subseteq> A}| \<le>o |afset ` {as. set as \<subseteq> A}|"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 428 | apply(rule card_of_mono1) unfolding fset_def2_raw apply auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 429 | apply (rule image_eqI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 430 | by (auto simp: afset_rfset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 431 |   also have "|afset ` {as. set as \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|" using card_of_image .
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 432 |   also have "|{as. set as \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd)
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 433 |   finally show "|{x. fset x \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 434 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 435 | fix A B1 B2 f1 f2 p1 p2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 436 | assume wp: "wpull A B1 B2 f1 f2 p1 p2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 437 | hence "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 438 | by (rule wpull_image) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 439 |   show "wpull {x. fset x \<subseteq> A} {x. fset x \<subseteq> B1} {x. fset x \<subseteq> B2}
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 440 | (map_fset f1) (map_fset f2) (map_fset p1) (map_fset p2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 441 | unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 442 | fix y1 y2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 443 | assume Y1: "fset y1 \<subseteq> B1" and Y2: "fset y2 \<subseteq> B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 444 | assume "map_fset f1 y1 = map_fset f2 y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 445 | hence EQ: "f1 ` (fset y1) = f2 ` (fset y2)" unfolding map_fset_def2_raw | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 446 | unfolding afset_set set_map fset_def2_raw . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 447 | with Y1 Y2 obtain X where X: "X \<subseteq> A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 448 | and Y1: "p1 ` X = fset y1" and Y2: "p2 ` X = fset y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 449 | using wpull_image[OF wp] unfolding wpull_def Pow_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 450 | unfolding Bex_def mem_Collect_eq apply - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 451 | apply(erule allE[of _ "fset y1"], erule allE[of _ "fset y2"]) by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 452 | have "\<forall> y1' \<in> fset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 453 | then obtain q1 where q1: "\<forall> y1' \<in> fset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 454 | have "\<forall> y2' \<in> fset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 455 | then obtain q2 where q2: "\<forall> y2' \<in> fset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 456 | def X' \<equiv> "q1 ` (fset y1) \<union> q2 ` (fset y2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 457 | have X': "X' \<subseteq> A" and Y1: "p1 ` X' = fset y1" and Y2: "p2 ` X' = fset y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 458 | using X Y1 Y2 q1 q2 unfolding X'_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 459 | have fX': "finite X'" unfolding X'_def by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 460 | then obtain x where X'eq: "X' = fset x" by (auto dest: finite_ex_fset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 461 | show "\<exists>x. fset x \<subseteq> A \<and> map_fset p1 x = y1 \<and> map_fset p2 x = y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 462 | apply(intro exI[of _ "x"]) using X' Y1 Y2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 463 | unfolding X'eq map_fset_def2_raw fset_def2_raw set_map[symmetric] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 464 | afset_set[symmetric] afset_rfset by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 465 | qed | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 466 | next | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 467 | fix R | 
| 49507 | 468 |   show "{p. fset_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
 | 
| 49463 | 469 |         (Gr {x. fset x \<subseteq> R} (map_fset fst))\<inverse> O Gr {x. fset x \<subseteq> R} (map_fset snd)"
 | 
| 49507 | 470 | unfolding fset_rel_def fset_rel_aux by simp | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 471 | qed auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 472 | |
| 49877 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 473 | lemma fset_rel_fset: "set_rel \<chi> (fset A1) (fset A2) = fset_rel \<chi> A1 A2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 474 | unfolding fset_rel_def set_rel_def by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 475 | |
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 476 | (* Countable sets *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 477 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 478 | lemma card_of_countable_sets_range: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 479 | fixes A :: "'a set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 480 | shows "|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |{f::nat \<Rightarrow> 'a. range f \<subseteq> A}|"
 | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50027diff
changeset | 481 | apply(rule card_of_ordLeqI[of from_nat_into]) using inj_on_from_nat_into | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 482 | unfolding inj_on_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 483 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 484 | lemma card_of_countable_sets_Func: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 485 | "|{X. X \<subseteq> A \<and> countable X \<and> X \<noteq> {}}| \<le>o |A| ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 486 | using card_of_countable_sets_range card_of_Func_UNIV[THEN ordIso_symmetric] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 487 | unfolding cexp_def Field_natLeq Field_card_of | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 488 | by (rule ordLeq_ordIso_trans) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 489 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 490 | lemma ordLeq_countable_subsets: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 491 | "|A| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 492 | apply (rule card_of_ordLeqI[of "\<lambda> a. {a}"]) unfolding inj_on_def by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 493 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 494 | lemma finite_countable_subset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 495 | "finite {X. X \<subseteq> A \<and> countable X} \<longleftrightarrow> finite A"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 496 | apply default | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 497 | apply (erule contrapos_pp) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 498 | apply (rule card_of_ordLeq_infinite) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 499 | apply (rule ordLeq_countable_subsets) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 500 | apply assumption | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 501 | apply (rule finite_Collect_conjI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 502 | apply (rule disjI1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 503 | by (erule finite_Collect_subsets) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 504 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 505 | lemma card_of_countable_sets: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 506 | "|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 507 | (is "|?L| \<le>o _") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 508 | proof(cases "finite A") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 509 | let ?R = "Func (UNIV::nat set) (A <+> (UNIV::bool set))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 510 | case True hence "finite ?L" by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 511 | moreover have "infinite ?R" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 512 | apply(rule infinite_Func[of _ "Inr True" "Inr False"]) by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 513 | ultimately show ?thesis unfolding cexp_def csum_def ctwo_def Field_natLeq Field_card_of | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 514 | apply(intro ordLess_imp_ordLeq) by (rule finite_ordLess_infinite2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 515 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 516 | case False | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 517 |   hence "|{X. X \<subseteq> A \<and> countable X}| =o |{X. X \<subseteq> A \<and> countable X} - {{}}|"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 518 | by (intro card_of_infinite_diff_finitte finite.emptyI finite.insertI ordIso_symmetric) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 519 | (unfold finite_countable_subset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 520 |   also have "|{X. X \<subseteq> A \<and> countable X} - {{}}| \<le>o |A| ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 521 | using card_of_countable_sets_Func[of A] unfolding set_diff_eq by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 522 | also have "|A| ^c natLeq \<le>o ( |A| +c ctwo) ^c natLeq" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 523 | apply(rule cexp_mono1_cone_ordLeq) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 524 | apply(rule ordLeq_csum1, rule card_of_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 525 | apply (rule cone_ordLeq_cexp) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 526 | apply (rule cone_ordLeq_Cnotzero) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 527 | using csum_Cnotzero2 ctwo_Cnotzero apply blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 528 | by (rule natLeq_Card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 529 | finally show ?thesis . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 530 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 531 | |
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 532 | lemma rcset_to_rcset: "countable A \<Longrightarrow> rcset (the_inv rcset A) = A" | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 533 | apply (rule f_the_inv_into_f) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 534 | unfolding inj_on_def rcset_inj using rcset_surj by auto | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 535 | |
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 536 | lemma Collect_Int_Times: | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 537 | "{(x, y). R x y} \<inter> A \<times> B = {(x, y). R x y \<and> x \<in> A \<and> y \<in> B}"
 | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 538 | by auto | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 539 | |
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 540 | lemma rcset_natural': "rcset (cIm f x) = f ` rcset x" | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 541 | unfolding cIm_def[abs_def] by simp | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 542 | |
| 49507 | 543 | definition cset_rel :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a cset \<Rightarrow> 'b cset \<Rightarrow> bool" where
 | 
| 544 | "cset_rel R a b \<longleftrightarrow> | |
| 49463 | 545 | (\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and> | 
| 546 | (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t)" | |
| 547 | ||
| 49507 | 548 | lemma cset_rel_aux: | 
| 49463 | 549 | "(\<forall>t \<in> rcset a. \<exists>u \<in> rcset b. R t u) \<and> (\<forall>t \<in> rcset b. \<exists>u \<in> rcset a. R u t) \<longleftrightarrow> | 
| 550 |  (a, b) \<in> (Gr {x. rcset x \<subseteq> {(a, b). R a b}} (cIm fst))\<inverse> O
 | |
| 551 |           Gr {x. rcset x \<subseteq> {(a, b). R a b}} (cIm snd)" (is "?L = ?R")
 | |
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 552 | proof | 
| 49463 | 553 | assume ?L | 
| 554 | def R' \<equiv> "the_inv rcset (Collect (split R) \<inter> (rcset a \<times> rcset b))" | |
| 555 | (is "the_inv rcset ?L'") | |
| 556 | have "countable ?L'" by auto | |
| 557 | hence *: "rcset R' = ?L'" unfolding R'_def using fset_to_fset by (intro rcset_to_rcset) | |
| 558 | show ?R unfolding Gr_def relcomp_unfold converse_unfold | |
| 559 | proof (intro CollectI prod_caseI exI conjI) | |
| 560 |     have "rcset a = fst ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?A")
 | |
| 561 | using conjunct1[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times) | |
| 562 | hence "a = acset ?A" by (metis acset_rcset) | |
| 563 | thus "(R', a) = (R', cIm fst R')" unfolding cIm_def * by auto | |
| 564 |     have "rcset b = snd ` ({(x, y). R x y} \<inter> rcset a \<times> rcset b)" (is "_ = ?B")
 | |
| 565 | using conjunct2[OF `?L`] unfolding image_def by (auto simp add: Collect_Int_Times) | |
| 566 | hence "b = acset ?B" by (metis acset_rcset) | |
| 567 | thus "(R', b) = (R', cIm snd R')" unfolding cIm_def * by auto | |
| 568 | qed (auto simp add: *) | |
| 569 | next | |
| 570 | assume ?R thus ?L unfolding Gr_def relcomp_unfold converse_unfold | |
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 571 | apply (simp add: subset_eq Ball_def) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 572 | apply (rule conjI) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 573 | apply (clarsimp, metis (lifting, no_types) rcset_natural' image_iff surjective_pairing) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 574 | apply (clarsimp) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 575 | by (metis Domain.intros Range.simps rcset_natural' fst_eq_Domain snd_eq_Range) | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 576 | qed | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 577 | |
| 49507 | 578 | bnf_def cIm [rcset] "\<lambda>_::'a cset. natLeq" ["cEmp"] cset_rel | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 579 | proof - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 580 | show "cIm id = id" unfolding cIm_def[abs_def] id_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 581 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 582 | fix f g show "cIm (g \<circ> f) = cIm g \<circ> cIm f" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 583 | unfolding cIm_def[abs_def] apply(rule ext) unfolding comp_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 584 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 585 | fix C f g assume eq: "\<And>a. a \<in> rcset C \<Longrightarrow> f a = g a" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 586 | thus "cIm f C = cIm g C" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 587 | unfolding cIm_def[abs_def] unfolding image_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 588 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 589 | fix f show "rcset \<circ> cIm f = op ` f \<circ> rcset" unfolding cIm_def[abs_def] by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 590 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 591 | show "card_order natLeq" by (rule natLeq_card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 592 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 593 | show "cinfinite natLeq" by (rule natLeq_cinfinite) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 594 | next | 
| 50144 
885deccc264e
renamed BNF/Countable_Set to Countable_Type and moved its generic stuff to Library/Countable_Set
 hoelzl parents: 
50027diff
changeset | 595 | fix C show "|rcset C| \<le>o natLeq" using rcset unfolding countable_card_le_natLeq . | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 596 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 597 | fix A :: "'a set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 598 |   have "|{Z. rcset Z \<subseteq> A}| \<le>o |acset ` {X. X \<subseteq> A \<and> countable X}|"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 599 | apply(rule card_of_mono1) unfolding Pow_def image_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 600 | proof (rule Collect_mono, clarsimp) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 601 | fix x | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 602 | assume "rcset x \<subseteq> A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 603 | hence "rcset x \<subseteq> A \<and> countable (rcset x) \<and> x = acset (rcset x)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 604 | using acset_rcset[of x] rcset[of x] by force | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 605 | thus "\<exists>y \<subseteq> A. countable y \<and> x = acset y" by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 606 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 607 |   also have "|acset ` {X. X \<subseteq> A \<and> countable X}| \<le>o |{X. X \<subseteq> A \<and> countable X}|"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 608 | using card_of_image . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 609 |   also have "|{X. X \<subseteq> A \<and> countable X}| \<le>o ( |A| +c ctwo) ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 610 | using card_of_countable_sets . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 611 |   finally show "|{Z. rcset Z \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 612 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 613 | fix A B1 B2 f1 f2 p1 p2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 614 | assume wp: "wpull A B1 B2 f1 f2 p1 p2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 615 |   show "wpull {x. rcset x \<subseteq> A} {x. rcset x \<subseteq> B1} {x. rcset x \<subseteq> B2}
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 616 | (cIm f1) (cIm f2) (cIm p1) (cIm p2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 617 | unfolding wpull_def proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 618 | fix y1 y2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 619 | assume Y1: "rcset y1 \<subseteq> B1" and Y2: "rcset y2 \<subseteq> B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 620 | assume "cIm f1 y1 = cIm f2 y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 621 | hence EQ: "f1 ` (rcset y1) = f2 ` (rcset y2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 622 | unfolding cIm_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 623 | with Y1 Y2 obtain X where X: "X \<subseteq> A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 624 | and Y1: "p1 ` X = rcset y1" and Y2: "p2 ` X = rcset y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 625 | using wpull_image[OF wp] unfolding wpull_def Pow_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 626 | unfolding Bex_def mem_Collect_eq apply - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 627 | apply(erule allE[of _ "rcset y1"], erule allE[of _ "rcset y2"]) by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 628 | have "\<forall> y1' \<in> rcset y1. \<exists> x. x \<in> X \<and> y1' = p1 x" using Y1 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 629 | then obtain q1 where q1: "\<forall> y1' \<in> rcset y1. q1 y1' \<in> X \<and> y1' = p1 (q1 y1')" by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 630 | have "\<forall> y2' \<in> rcset y2. \<exists> x. x \<in> X \<and> y2' = p2 x" using Y2 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 631 | then obtain q2 where q2: "\<forall> y2' \<in> rcset y2. q2 y2' \<in> X \<and> y2' = p2 (q2 y2')" by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 632 | def X' \<equiv> "q1 ` (rcset y1) \<union> q2 ` (rcset y2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 633 | have X': "X' \<subseteq> A" and Y1: "p1 ` X' = rcset y1" and Y2: "p2 ` X' = rcset y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 634 | using X Y1 Y2 q1 q2 unfolding X'_def by fast+ | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 635 | have fX': "countable X'" unfolding X'_def by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 636 | then obtain x where X'eq: "X' = rcset x" by (metis rcset_acset) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 637 |     show "\<exists>x\<in>{x. rcset x \<subseteq> A}. cIm p1 x = y1 \<and> cIm p2 x = y2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 638 | apply(intro bexI[of _ "x"]) using X' Y1 Y2 unfolding X'eq cIm_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 639 | qed | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 640 | next | 
| 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 641 | fix R | 
| 49507 | 642 |   show "{p. cset_rel (\<lambda>x y. (x, y) \<in> R) (fst p) (snd p)} =
 | 
| 49463 | 643 |         (Gr {x. rcset x \<subseteq> R} (cIm fst))\<inverse> O Gr {x. rcset x \<subseteq> R} (cIm snd)"
 | 
| 49507 | 644 | unfolding cset_rel_def cset_rel_aux by simp | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 645 | qed (unfold cEmp_def, auto) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 646 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 647 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 648 | (* Multisets *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 649 | |
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 650 | (* The cardinal of a mutiset: this, and the following basic lemmas about it, | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 651 | should eventually go into Multiset.thy *) | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 652 | definition "mcard M \<equiv> setsum (count M) {a. count M a > 0}"
 | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 653 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 654 | lemma mcard_emp[simp]: "mcard {#} = 0"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 655 | unfolding mcard_def by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 656 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 657 | lemma mcard_emp_iff[simp]: "mcard M = 0 \<longleftrightarrow> M = {#}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 658 | unfolding mcard_def apply safe | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 659 | apply simp_all | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 660 | by (metis multi_count_eq zero_multiset.rep_eq) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 661 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 662 | lemma mcard_singl[simp]: "mcard {#a#} = Suc 0"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 663 | unfolding mcard_def by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 664 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 665 | lemma mcard_Plus[simp]: "mcard (M + N) = mcard M + mcard N" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 666 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 667 |   have "setsum (count M) {a. 0 < count M a + count N a} =
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 668 |         setsum (count M) {a. a \<in># M}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 669 | apply(rule setsum_mono_zero_cong_right) by auto | 
| 49461 
de07eecb2664
adapting "More_BNFs" to new relators/predicators
 blanchet parents: 
49440diff
changeset | 670 | moreover | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 671 |   have "setsum (count N) {a. 0 < count M a + count N a} =
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 672 |         setsum (count N) {a. a \<in># N}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 673 | apply(rule setsum_mono_zero_cong_right) by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 674 | ultimately show ?thesis | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 675 | unfolding mcard_def count_union[THEN ext] comm_monoid_add_class.setsum.F_fun_f by simp | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 676 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 677 | |
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 678 | lemma setsum_gt_0_iff: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 679 | fixes f :: "'a \<Rightarrow> nat" assumes "finite A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 680 | shows "setsum f A > 0 \<longleftrightarrow> (\<exists> a \<in> A. f a > 0)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 681 | (is "?L \<longleftrightarrow> ?R") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 682 | proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 683 | have "?L \<longleftrightarrow> \<not> setsum f A = 0" by fast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 684 | also have "... \<longleftrightarrow> (\<exists> a \<in> A. f a \<noteq> 0)" using assms by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 685 | also have "... \<longleftrightarrow> ?R" by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 686 | finally show ?thesis . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 687 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 688 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 689 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 690 | definition mmap :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a \<Rightarrow> nat) \<Rightarrow> 'b \<Rightarrow> nat" where
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 691 | "mmap h f b = setsum f {a. h a = b \<and> f a > 0}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 692 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 693 | lemma mmap_id: "mmap id = id" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 694 | proof (rule ext)+ | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 695 | fix f a show "mmap id f a = id f a" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 696 | proof(cases "f a = 0") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 697 | case False | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 698 |     hence 1: "{aa. aa = a \<and> 0 < f aa} = {a}" by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 699 | show ?thesis by (simp add: mmap_def id_apply 1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 700 | qed(unfold mmap_def, auto) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 701 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 702 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 703 | lemma inj_on_setsum_inv: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 704 | assumes f: "f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 705 | and 1: "(0::nat) < setsum f {a. h a = b' \<and> 0 < f a}" (is "0 < setsum f ?A'")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 706 | and 2: "{a. h a = b \<and> 0 < f a} = {a. h a = b' \<and> 0 < f a}" (is "?A = ?A'")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 707 | shows "b = b'" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 708 | proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 709 | have "finite ?A'" using f unfolding multiset_def by auto | 
| 50027 
7747a9f4c358
adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
 bulwahn parents: 
49878diff
changeset | 710 |   hence "?A' \<noteq> {}" using 1 by (auto simp add: setsum_gt_0_iff)
 | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 711 | thus ?thesis using 2 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 712 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 713 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 714 | lemma mmap_comp: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 715 | fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 716 | assumes f: "f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 717 | shows "mmap (h2 o h1) f = (mmap h2 o mmap h1) f" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 718 | unfolding mmap_def[abs_def] comp_def proof(rule ext)+ | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 719 | fix c :: 'c | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 720 |   let ?A = "{a. h2 (h1 a) = c \<and> 0 < f a}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 721 |   let ?As = "\<lambda> b. {a. h1 a = b \<and> 0 < f a}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 722 |   let ?B = "{b. h2 b = c \<and> 0 < setsum f (?As b)}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 723 |   have 0: "{?As b | b.  b \<in> ?B} = ?As ` ?B" by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 724 | have "\<And> b. finite (?As b)" using f unfolding multiset_def by simp | 
| 50027 
7747a9f4c358
adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
 bulwahn parents: 
49878diff
changeset | 725 |   hence "?B = {b. h2 b = c \<and> ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
 | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 726 |   hence A: "?A = \<Union> {?As b | b.  b \<in> ?B}" by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 727 |   have "setsum f ?A = setsum (setsum f) {?As b | b.  b \<in> ?B}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 728 | unfolding A apply(rule setsum_Union_disjoint) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 729 | using f unfolding multiset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 730 | also have "... = setsum (setsum f) (?As ` ?B)" unfolding 0 .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 731 | also have "... = setsum (setsum f o ?As) ?B" apply(rule setsum_reindex) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 732 | unfolding inj_on_def apply auto using inj_on_setsum_inv[OF f, of h1] by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 733 | also have "... = setsum (\<lambda> b. setsum f (?As b)) ?B" unfolding comp_def .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 734 | finally show "setsum f ?A = setsum (\<lambda> b. setsum f (?As b)) ?B" . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 735 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 736 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 737 | lemma mmap_comp1: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 738 | fixes h1 :: "'a \<Rightarrow> 'b" and h2 :: "'b \<Rightarrow> 'c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 739 | assumes "f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 740 | shows "mmap (\<lambda> a. h2 (h1 a)) f = mmap h2 (mmap h1 f)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 741 | using mmap_comp[OF assms] unfolding comp_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 742 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 743 | lemma mmap: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 744 | assumes "f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 745 | shows "mmap h f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 746 | using assms unfolding mmap_def[abs_def] multiset_def proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 747 |   assume fin: "finite {a. 0 < f a}"  (is "finite ?A")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 748 |   show "finite {b. 0 < setsum f {a. h a = b \<and> 0 < f a}}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 749 |   (is "finite {b. 0 < setsum f (?As b)}")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 750 |   proof- let ?B = "{b. 0 < setsum f (?As b)}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 751 | have "\<And> b. finite (?As b)" using assms unfolding multiset_def by simp | 
| 50027 
7747a9f4c358
adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
 bulwahn parents: 
49878diff
changeset | 752 |     hence B: "?B = {b. ?As b \<noteq> {}}" by (auto simp add: setsum_gt_0_iff)
 | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 753 | hence "?B \<subseteq> h ` ?A" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 754 | thus ?thesis using finite_surj[OF fin] by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 755 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 756 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 757 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 758 | lemma mmap_cong: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 759 | assumes "\<And>a. a \<in># M \<Longrightarrow> f a = g a" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 760 | shows "mmap f (count M) = mmap g (count M)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 761 | using assms unfolding mmap_def[abs_def] by (intro ext, intro setsum_cong) auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 762 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 763 | abbreviation supp where "supp f \<equiv> {a. f a > 0}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 764 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 765 | lemma mmap_image_comp: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 766 | assumes f: "f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 767 | shows "(supp o mmap h) f = (image h o supp) f" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 768 | unfolding mmap_def[abs_def] comp_def proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 769 |   have "\<And> b. finite {a. h a = b \<and> 0 < f a}" (is "\<And> b. finite (?As b)")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 770 | using f unfolding multiset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 771 |   thus "{b. 0 < setsum f (?As b)} = h ` {a. 0 < f a}"
 | 
| 50027 
7747a9f4c358
adjusting proofs as the set_comprehension_pointfree simproc breaks some existing proofs
 bulwahn parents: 
49878diff
changeset | 772 | by (auto simp add: setsum_gt_0_iff) | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 773 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 774 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 775 | lemma mmap_image: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 776 | assumes f: "f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 777 | shows "supp (mmap h f) = h ` (supp f)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 778 | using mmap_image_comp[OF assms] unfolding comp_def . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 779 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 780 | lemma set_of_Abs_multiset: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 781 | assumes f: "f \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 782 | shows "set_of (Abs_multiset f) = supp f" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 783 | using assms unfolding set_of_def by (auto simp: Abs_multiset_inverse) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 784 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 785 | lemma supp_count: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 786 | "supp (count M) = set_of M" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 787 | using assms unfolding set_of_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 788 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 789 | lemma multiset_of_surj: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 790 | "multiset_of ` {as. set as \<subseteq> A} = {M. set_of M \<subseteq> A}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 791 | proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 792 | fix M assume M: "set_of M \<subseteq> A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 793 | obtain as where eq: "M = multiset_of as" using surj_multiset_of unfolding surj_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 794 | hence "set as \<subseteq> A" using M by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 795 |   thus "M \<in> multiset_of ` {as. set as \<subseteq> A}" using eq by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 796 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 797 | show "\<And>x xa xb. \<lbrakk>set xa \<subseteq> A; xb \<in> set_of (multiset_of xa)\<rbrakk> \<Longrightarrow> xb \<in> A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 798 | by (erule set_mp) (unfold set_of_multiset_of) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 799 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 800 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 801 | lemma card_of_set_of: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 802 | "|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 803 | apply(rule card_of_ordLeqI2[of _ multiset_of]) using multiset_of_surj by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 804 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 805 | lemma nat_sum_induct: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 806 | assumes "\<And>n1 n2. (\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow> phi m1 m2) \<Longrightarrow> phi n1 n2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 807 | shows "phi (n1::nat) (n2::nat)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 808 | proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 809 | let ?chi = "\<lambda> n1n2 :: nat * nat. phi (fst n1n2) (snd n1n2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 810 | have "?chi (n1,n2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 811 | apply(induct rule: measure_induct[of "\<lambda> n1n2. fst n1n2 + snd n1n2" ?chi]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 812 | using assms by (metis fstI sndI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 813 | thus ?thesis by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 814 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 815 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 816 | lemma matrix_count: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 817 | fixes ct1 ct2 :: "nat \<Rightarrow> nat" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 818 | assumes "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 819 | shows | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 820 | "\<exists> ct. (\<forall> i1 \<le> n1. setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ct1 i1) \<and>
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 821 |        (\<forall> i2 \<le> n2. setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ct2 i2)"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 822 | (is "?phi ct1 ct2 n1 n2") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 823 | proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 824 | have "\<forall> ct1 ct2 :: nat \<Rightarrow> nat. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 825 |         setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 826 | proof(induct rule: nat_sum_induct[of | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 827 | "\<lambda> n1 n2. \<forall> ct1 ct2 :: nat \<Rightarrow> nat. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 828 |      setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2} \<longrightarrow> ?phi ct1 ct2 n1 n2"],
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 829 | clarify) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 830 | fix n1 n2 :: nat and ct1 ct2 :: "nat \<Rightarrow> nat" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 831 | assume IH: "\<And> m1 m2. m1 + m2 < n1 + n2 \<Longrightarrow> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 832 | \<forall> dt1 dt2 :: nat \<Rightarrow> nat. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 833 |                 setsum dt1 {..<Suc m1} = setsum dt2 {..<Suc m2} \<longrightarrow> ?phi dt1 dt2 m1 m2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 834 |   and ss: "setsum ct1 {..<Suc n1} = setsum ct2 {..<Suc n2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 835 | show "?phi ct1 ct2 n1 n2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 836 | proof(cases n1) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 837 | case 0 note n1 = 0 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 838 | show ?thesis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 839 | proof(cases n2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 840 | case 0 note n2 = 0 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 841 | let ?ct = "\<lambda> i1 i2. ct2 0" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 842 | show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 843 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 844 | case (Suc m2) note n2 = Suc | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 845 | let ?ct = "\<lambda> i1 i2. ct2 i2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 846 | show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 847 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 848 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 849 | case (Suc m1) note n1 = Suc | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 850 | show ?thesis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 851 | proof(cases n2) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 852 | case 0 note n2 = 0 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 853 | let ?ct = "\<lambda> i1 i2. ct1 i1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 854 | show ?thesis apply(rule exI[of _ ?ct]) using n1 n2 ss by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 855 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 856 | case (Suc m2) note n2 = Suc | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 857 | show ?thesis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 858 | proof(cases "ct1 n1 \<le> ct2 n2") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 859 | case True | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 860 | def dt2 \<equiv> "\<lambda> i2. if i2 = n2 then ct2 i2 - ct1 n1 else ct2 i2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 861 |         have "setsum ct1 {..<Suc m1} = setsum dt2 {..<Suc n2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 862 | unfolding dt2_def using ss n1 True by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 863 | hence "?phi ct1 dt2 m1 n2" using IH[of m1 n2] n1 by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 864 | then obtain dt where | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 865 |         1: "\<And> i1. i1 \<le> m1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc n2} = ct1 i1" and
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 866 |         2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc m1} = dt2 i2" by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 867 | let ?ct = "\<lambda> i1 i2. if i1 = n1 then (if i2 = n2 then ct1 n1 else 0) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 868 | else dt i1 i2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 869 | show ?thesis apply(rule exI[of _ ?ct]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 870 | using n1 n2 1 2 True unfolding dt2_def by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 871 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 872 | case False | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 873 | hence False: "ct2 n2 < ct1 n1" by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 874 | def dt1 \<equiv> "\<lambda> i1. if i1 = n1 then ct1 i1 - ct2 n2 else ct1 i1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 875 |         have "setsum dt1 {..<Suc n1} = setsum ct2 {..<Suc m2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 876 | unfolding dt1_def using ss n2 False by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 877 | hence "?phi dt1 ct2 n1 m2" using IH[of n1 m2] n2 by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 878 | then obtain dt where | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 879 |         1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. dt i1 i2) {..<Suc m2} = dt1 i1" and
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 880 |         2: "\<And> i2. i2 \<le> m2 \<Longrightarrow> setsum (\<lambda> i1. dt i1 i2) {..<Suc n1} = ct2 i2" by force
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 881 | let ?ct = "\<lambda> i1 i2. if i2 = n2 then (if i1 = n1 then ct2 n2 else 0) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 882 | else dt i1 i2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 883 | show ?thesis apply(rule exI[of _ ?ct]) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 884 | using n1 n2 1 2 False unfolding dt1_def by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 885 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 886 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 887 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 888 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 889 | thus ?thesis using assms by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 890 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 891 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 892 | definition | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 893 | "inj2 u B1 B2 \<equiv> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 894 |  \<forall> b1 b1' b2 b2'. {b1,b1'} \<subseteq> B1 \<and> {b2,b2'} \<subseteq> B2 \<and> u b1 b2 = u b1' b2'
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 895 | \<longrightarrow> b1 = b1' \<and> b2 = b2'" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 896 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 897 | lemma matrix_setsum_finite: | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 898 | assumes B1: "B1 \<noteq> {}" "finite B1" and B2: "B2 \<noteq> {}" "finite B2" and u: "inj2 u B1 B2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 899 | and ss: "setsum N1 B1 = setsum N2 B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 900 | shows "\<exists> M :: 'a \<Rightarrow> nat. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 901 | (\<forall> b1 \<in> B1. setsum (\<lambda> b2. M (u b1 b2)) B2 = N1 b1) \<and> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 902 | (\<forall> b2 \<in> B2. setsum (\<lambda> b1. M (u b1 b2)) B1 = N2 b2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 903 | proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 904 | obtain n1 where "card B1 = Suc n1" using B1 by (metis card_insert finite.simps) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 905 |   then obtain e1 where e1: "bij_betw e1 {..<Suc n1} B1"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 906 | using ex_bij_betw_finite_nat[OF B1(2)] by (metis atLeast0LessThan bij_betw_the_inv_into) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 907 |   hence e1_inj: "inj_on e1 {..<Suc n1}" and e1_surj: "e1 ` {..<Suc n1} = B1"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 908 | unfolding bij_betw_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 909 |   def f1 \<equiv> "inv_into {..<Suc n1} e1"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 910 |   have f1: "bij_betw f1 B1 {..<Suc n1}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 911 | and f1e1[simp]: "\<And> i1. i1 < Suc n1 \<Longrightarrow> f1 (e1 i1) = i1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 912 | and e1f1[simp]: "\<And> b1. b1 \<in> B1 \<Longrightarrow> e1 (f1 b1) = b1" unfolding f1_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 913 | apply (metis bij_betw_inv_into e1, metis bij_betw_inv_into_left e1 lessThan_iff) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 914 | by (metis e1_surj f_inv_into_f) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 915 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 916 | obtain n2 where "card B2 = Suc n2" using B2 by (metis card_insert finite.simps) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 917 |   then obtain e2 where e2: "bij_betw e2 {..<Suc n2} B2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 918 | using ex_bij_betw_finite_nat[OF B2(2)] by (metis atLeast0LessThan bij_betw_the_inv_into) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 919 |   hence e2_inj: "inj_on e2 {..<Suc n2}" and e2_surj: "e2 ` {..<Suc n2} = B2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 920 | unfolding bij_betw_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 921 |   def f2 \<equiv> "inv_into {..<Suc n2} e2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 922 |   have f2: "bij_betw f2 B2 {..<Suc n2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 923 | and f2e2[simp]: "\<And> i2. i2 < Suc n2 \<Longrightarrow> f2 (e2 i2) = i2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 924 | and e2f2[simp]: "\<And> b2. b2 \<in> B2 \<Longrightarrow> e2 (f2 b2) = b2" unfolding f2_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 925 | apply (metis bij_betw_inv_into e2, metis bij_betw_inv_into_left e2 lessThan_iff) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 926 | by (metis e2_surj f_inv_into_f) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 927 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 928 | let ?ct1 = "N1 o e1" let ?ct2 = "N2 o e2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 929 |   have ss: "setsum ?ct1 {..<Suc n1} = setsum ?ct2 {..<Suc n2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 930 | unfolding setsum_reindex[OF e1_inj, symmetric] setsum_reindex[OF e2_inj, symmetric] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 931 | e1_surj e2_surj using ss . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 932 | obtain ct where | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 933 |   ct1: "\<And> i1. i1 \<le> n1 \<Longrightarrow> setsum (\<lambda> i2. ct i1 i2) {..<Suc n2} = ?ct1 i1" and
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 934 |   ct2: "\<And> i2. i2 \<le> n2 \<Longrightarrow> setsum (\<lambda> i1. ct i1 i2) {..<Suc n1} = ?ct2 i2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 935 | using matrix_count[OF ss] by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 936 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 937 |   def A \<equiv> "{u b1 b2 | b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 938 | have "\<forall> a \<in> A. \<exists> b1b2 \<in> B1 <*> B2. u (fst b1b2) (snd b1b2) = a" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 939 | unfolding A_def Ball_def mem_Collect_eq by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 940 | then obtain h1h2 where h12: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 941 | "\<And>a. a \<in> A \<Longrightarrow> u (fst (h1h2 a)) (snd (h1h2 a)) = a \<and> h1h2 a \<in> B1 <*> B2" by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 942 | def h1 \<equiv> "fst o h1h2" def h2 \<equiv> "snd o h1h2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 943 | have h12[simp]: "\<And>a. a \<in> A \<Longrightarrow> u (h1 a) (h2 a) = a" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 944 | "\<And> a. a \<in> A \<Longrightarrow> h1 a \<in> B1" "\<And> a. a \<in> A \<Longrightarrow> h2 a \<in> B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 945 | using h12 unfolding h1_def h2_def by force+ | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 946 |   {fix b1 b2 assume b1: "b1 \<in> B1" and b2: "b2 \<in> B2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 947 | hence inA: "u b1 b2 \<in> A" unfolding A_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 948 | hence "u b1 b2 = u (h1 (u b1 b2)) (h2 (u b1 b2))" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 949 | moreover have "h1 (u b1 b2) \<in> B1" "h2 (u b1 b2) \<in> B2" using inA by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 950 | ultimately have "h1 (u b1 b2) = b1 \<and> h2 (u b1 b2) = b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 951 | using u b1 b2 unfolding inj2_def by fastforce | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 952 | } | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 953 | hence h1[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h1 (u b1 b2) = b1" and | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 954 | h2[simp]: "\<And> b1 b2. \<lbrakk>b1 \<in> B1; b2 \<in> B2\<rbrakk> \<Longrightarrow> h2 (u b1 b2) = b2" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 955 | def M \<equiv> "\<lambda> a. ct (f1 (h1 a)) (f2 (h2 a))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 956 | show ?thesis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 957 | apply(rule exI[of _ M]) proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 958 | fix b1 assume b1: "b1 \<in> B1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 959 | hence f1b1: "f1 b1 \<le> n1" using f1 unfolding bij_betw_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 960 | by (metis bij_betwE f1 lessThan_iff less_Suc_eq_le) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 961 | have "(\<Sum>b2\<in>B2. M (u b1 b2)) = (\<Sum>i2<Suc n2. ct (f1 b1) (f2 (e2 i2)))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 962 | unfolding e2_surj[symmetric] setsum_reindex[OF e2_inj] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 963 | unfolding M_def comp_def apply(intro setsum_cong) apply force | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 964 | by (metis e2_surj b1 h1 h2 imageI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 965 | also have "... = N1 b1" using b1 ct1[OF f1b1] by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 966 | finally show "(\<Sum>b2\<in>B2. M (u b1 b2)) = N1 b1" . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 967 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 968 | fix b2 assume b2: "b2 \<in> B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 969 | hence f2b2: "f2 b2 \<le> n2" using f2 unfolding bij_betw_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 970 | by (metis bij_betwE f2 lessThan_iff less_Suc_eq_le) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 971 | have "(\<Sum>b1\<in>B1. M (u b1 b2)) = (\<Sum>i1<Suc n1. ct (f1 (e1 i1)) (f2 b2))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 972 | unfolding e1_surj[symmetric] setsum_reindex[OF e1_inj] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 973 | unfolding M_def comp_def apply(intro setsum_cong) apply force | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 974 | by (metis e1_surj b2 h1 h2 imageI) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 975 | also have "... = N2 b2" using b2 ct2[OF f2b2] by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 976 | finally show "(\<Sum>b1\<in>B1. M (u b1 b2)) = N2 b2" . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 977 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 978 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 979 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 980 | lemma supp_vimage_mmap: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 981 | assumes "M \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 982 | shows "supp M \<subseteq> f -` (supp (mmap f M))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 983 | using assms by (auto simp: mmap_image) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 984 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 985 | lemma mmap_ge_0: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 986 | assumes "M \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 987 | shows "0 < mmap f M b \<longleftrightarrow> (\<exists>a. 0 < M a \<and> f a = b)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 988 | proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 989 |   have f: "finite {a. f a = b \<and> 0 < M a}" using assms unfolding multiset_def by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 990 | show ?thesis unfolding mmap_def setsum_gt_0_iff[OF f] by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 991 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 992 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 993 | lemma finite_twosets: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 994 | assumes "finite B1" and "finite B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 995 | shows "finite {u b1 b2 |b1 b2. b1 \<in> B1 \<and> b2 \<in> B2}"  (is "finite ?A")
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 996 | proof- | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 997 | have A: "?A = (\<lambda> b1b2. u (fst b1b2) (snd b1b2)) ` (B1 <*> B2)" by force | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 998 | show ?thesis unfolding A using finite_cartesian_product[OF assms] by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 999 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1000 | |
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1001 | lemma wp_mmap: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1002 | fixes A :: "'a set" and B1 :: "'b1 set" and B2 :: "'b2 set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1003 | assumes wp: "wpull A B1 B2 f1 f2 p1 p2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1004 | shows | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1005 | "wpull {M. M \<in> multiset \<and> supp M \<subseteq> A}
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1006 |        {N1. N1 \<in> multiset \<and> supp N1 \<subseteq> B1} {N2. N2 \<in> multiset \<and> supp N2 \<subseteq> B2}
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1007 | (mmap f1) (mmap f2) (mmap p1) (mmap p2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1008 | unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1009 | fix N1 :: "'b1 \<Rightarrow> nat" and N2 :: "'b2 \<Rightarrow> nat" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1010 | assume mmap': "mmap f1 N1 = mmap f2 N2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1011 | and N1[simp]: "N1 \<in> multiset" "supp N1 \<subseteq> B1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1012 | and N2[simp]: "N2 \<in> multiset" "supp N2 \<subseteq> B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1013 | have mN1[simp]: "mmap f1 N1 \<in> multiset" using N1 by (auto simp: mmap) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1014 | have mN2[simp]: "mmap f2 N2 \<in> multiset" using N2 by (auto simp: mmap) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1015 | def P \<equiv> "mmap f1 N1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1016 | have P1: "P = mmap f1 N1" and P2: "P = mmap f2 N2" unfolding P_def using mmap' by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1017 | note P = P1 P2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1018 | have P_mult[simp]: "P \<in> multiset" unfolding P_def using N1 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1019 | have fin_N1[simp]: "finite (supp N1)" using N1(1) unfolding multiset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1020 | have fin_N2[simp]: "finite (supp N2)" using N2(1) unfolding multiset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1021 | have fin_P[simp]: "finite (supp P)" using P_mult unfolding multiset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1022 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1023 |   def set1 \<equiv> "\<lambda> c. {b1 \<in> supp N1. f1 b1 = c}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1024 | have set1[simp]: "\<And> c b1. b1 \<in> set1 c \<Longrightarrow> f1 b1 = c" unfolding set1_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1025 | have fin_set1: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set1 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1026 | using N1(1) unfolding set1_def multiset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1027 |   have set1_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<noteq> {}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1028 | unfolding set1_def P1 mmap_ge_0[OF N1(1)] by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1029 | have supp_N1_set1: "supp N1 = (\<Union> c \<in> supp P. set1 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1030 | using supp_vimage_mmap[OF N1(1), of f1] unfolding set1_def P1 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1031 | hence set1_inclN1: "\<And>c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> supp N1" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1032 | hence set1_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set1 c \<subseteq> B1" using N1(2) by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1033 |   have set1_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set1 c \<inter> set1 c' = {}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1034 | unfolding set1_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1035 | have setsum_set1: "\<And> c. setsum N1 (set1 c) = P c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1036 | unfolding P1 set1_def mmap_def apply(rule setsum_cong) by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1037 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1038 |   def set2 \<equiv> "\<lambda> c. {b2 \<in> supp N2. f2 b2 = c}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1039 | have set2[simp]: "\<And> c b2. b2 \<in> set2 c \<Longrightarrow> f2 b2 = c" unfolding set2_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1040 | have fin_set2: "\<And> c. c \<in> supp P \<Longrightarrow> finite (set2 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1041 | using N2(1) unfolding set2_def multiset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1042 |   have set2_NE: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<noteq> {}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1043 | unfolding set2_def P2 mmap_ge_0[OF N2(1)] by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1044 | have supp_N2_set2: "supp N2 = (\<Union> c \<in> supp P. set2 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1045 | using supp_vimage_mmap[OF N2(1), of f2] unfolding set2_def P2 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1046 | hence set2_inclN2: "\<And>c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> supp N2" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1047 | hence set2_incl: "\<And> c. c \<in> supp P \<Longrightarrow> set2 c \<subseteq> B2" using N2(2) by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1048 |   have set2_disj: "\<And> c c'. c \<noteq> c' \<Longrightarrow> set2 c \<inter> set2 c' = {}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1049 | unfolding set2_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1050 | have setsum_set2: "\<And> c. setsum N2 (set2 c) = P c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1051 | unfolding P2 set2_def mmap_def apply(rule setsum_cong) by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1052 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1053 | have ss: "\<And> c. c \<in> supp P \<Longrightarrow> setsum N1 (set1 c) = setsum N2 (set2 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1054 | unfolding setsum_set1 setsum_set2 .. | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1055 | have "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c). | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1056 | \<exists> a \<in> A. p1 a = fst b1b2 \<and> p2 a = snd b1b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1057 | using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1058 | by simp (metis set1 set2 set_rev_mp) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1059 | then obtain uu where uu: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1060 | "\<forall> c \<in> supp P. \<forall> b1b2 \<in> (set1 c) \<times> (set2 c). | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1061 | uu c b1b2 \<in> A \<and> p1 (uu c b1b2) = fst b1b2 \<and> p2 (uu c b1b2) = snd b1b2" by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1062 | def u \<equiv> "\<lambda> c b1 b2. uu c (b1,b2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1063 | have u[simp]: | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1064 | "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1065 | "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p1 (u c b1 b2) = b1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1066 | "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> p2 (u c b1 b2) = b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1067 | using uu unfolding u_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1068 |   {fix c assume c: "c \<in> supp P"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1069 | have "inj2 (u c) (set1 c) (set2 c)" unfolding inj2_def proof clarify | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1070 | fix b1 b1' b2 b2' | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1071 |      assume "{b1, b1'} \<subseteq> set1 c" "{b2, b2'} \<subseteq> set2 c" and 0: "u c b1 b2 = u c b1' b2'"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1072 | hence "p1 (u c b1 b2) = b1 \<and> p2 (u c b1 b2) = b2 \<and> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1073 | p1 (u c b1' b2') = b1' \<and> p2 (u c b1' b2') = b2'" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1074 | using u(2)[OF c] u(3)[OF c] by simp metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1075 | thus "b1 = b1' \<and> b2 = b2'" using 0 by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1076 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1077 | } note inj = this | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1078 |   def sset \<equiv> "\<lambda> c. {u c b1 b2 | b1 b2. b1 \<in> set1 c \<and> b2 \<in> set2 c}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1079 | have fin_sset[simp]: "\<And> c. c \<in> supp P \<Longrightarrow> finite (sset c)" unfolding sset_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1080 | using fin_set1 fin_set2 finite_twosets by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1081 | have sset_A: "\<And> c. c \<in> supp P \<Longrightarrow> sset c \<subseteq> A" unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1082 |   {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1083 | then obtain b1 b2 where b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1084 | and a: "a = u c b1 b2" unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1085 | have "p1 a \<in> set1 c" and p2a: "p2 a \<in> set2 c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1086 | using ac a b1 b2 c u(2) u(3) by simp+ | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1087 | hence "u c (p1 a) (p2 a) = a" unfolding a using b1 b2 inj[OF c] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1088 | unfolding inj2_def by (metis c u(2) u(3)) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1089 | } note u_p12[simp] = this | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1090 |   {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1091 | hence "p1 a \<in> set1 c" unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1092 | }note p1[simp] = this | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1093 |   {fix c a assume c: "c \<in> supp P" and ac: "a \<in> sset c"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1094 | hence "p2 a \<in> set2 c" unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1095 | }note p2[simp] = this | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1096 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1097 |   {fix c assume c: "c \<in> supp P"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1098 | hence "\<exists> M. (\<forall> b1 \<in> set1 c. setsum (\<lambda> b2. M (u c b1 b2)) (set2 c) = N1 b1) \<and> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1099 | (\<forall> b2 \<in> set2 c. setsum (\<lambda> b1. M (u c b1 b2)) (set1 c) = N2 b2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1100 | unfolding sset_def | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1101 | using matrix_setsum_finite[OF set1_NE[OF c] fin_set1[OF c] | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1102 | set2_NE[OF c] fin_set2[OF c] inj[OF c] ss[OF c]] by auto | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1103 | } | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1104 | then obtain Ms where | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1105 | ss1: "\<And> c b1. \<lbrakk>c \<in> supp P; b1 \<in> set1 c\<rbrakk> \<Longrightarrow> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1106 | setsum (\<lambda> b2. Ms c (u c b1 b2)) (set2 c) = N1 b1" and | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1107 | ss2: "\<And> c b2. \<lbrakk>c \<in> supp P; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1108 | setsum (\<lambda> b1. Ms c (u c b1 b2)) (set1 c) = N2 b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1109 | by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1110 | def SET \<equiv> "\<Union> c \<in> supp P. sset c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1111 | have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1112 | have SET_A: "SET \<subseteq> A" unfolding SET_def using sset_A by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1113 | have u_SET[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> \<Longrightarrow> u c b1 b2 \<in> SET" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1114 | unfolding SET_def sset_def by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1115 |   {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p1a: "p1 a \<in> set1 c"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1116 | then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1117 | unfolding SET_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1118 | hence "p1 a \<in> set1 c'" unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1119 | hence eq: "c = c'" using p1a c c' set1_disj by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1120 | hence "a \<in> sset c" using ac' by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1121 | } note p1_rev = this | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1122 |   {fix c a assume c: "c \<in> supp P" and a: "a \<in> SET" and p2a: "p2 a \<in> set2 c"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1123 | then obtain c' where c': "c' \<in> supp P" and ac': "a \<in> sset c'" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1124 | unfolding SET_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1125 | hence "p2 a \<in> set2 c'" unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1126 | hence eq: "c = c'" using p2a c c' set2_disj by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1127 | hence "a \<in> sset c" using ac' by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1128 | } note p2_rev = this | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1129 | (* *) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1130 | have "\<forall> a \<in> SET. \<exists> c \<in> supp P. a \<in> sset c" unfolding SET_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1131 | then obtain h where h: "\<forall> a \<in> SET. h a \<in> supp P \<and> a \<in> sset (h a)" by metis | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1132 | have h_u[simp]: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1133 | \<Longrightarrow> h (u c b1 b2) = c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1134 | by (metis h p2 set2 u(3) u_SET) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1135 | have h_u1: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1136 | \<Longrightarrow> h (u c b1 b2) = f1 b1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1137 | using h unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1138 | have h_u2: "\<And> c b1 b2. \<lbrakk>c \<in> supp P; b1 \<in> set1 c; b2 \<in> set2 c\<rbrakk> | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1139 | \<Longrightarrow> h (u c b1 b2) = f2 b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1140 | using h unfolding sset_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1141 | def M \<equiv> "\<lambda> a. if a \<in> SET \<and> p1 a \<in> supp N1 \<and> p2 a \<in> supp N2 then Ms (h a) a else 0" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1142 | have sM: "supp M \<subseteq> SET" "supp M \<subseteq> p1 -` (supp N1)" "supp M \<subseteq> p2 -` (supp N2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1143 | unfolding M_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1144 | show "\<exists>M. (M \<in> multiset \<and> supp M \<subseteq> A) \<and> mmap p1 M = N1 \<and> mmap p2 M = N2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1145 | proof(rule exI[of _ M], safe) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1146 | show "M \<in> multiset" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1147 | unfolding multiset_def using finite_subset[OF sM(1) fin_SET] by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1148 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1149 | fix a assume "0 < M a" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1150 | thus "a \<in> A" unfolding M_def using SET_A by (cases "a \<in> SET") auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1151 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1152 | show "mmap p1 M = N1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1153 | unfolding mmap_def[abs_def] proof(rule ext) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1154 | fix b1 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1155 |       let ?K = "{a. p1 a = b1 \<and> 0 < M a}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1156 | show "setsum M ?K = N1 b1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1157 | proof(cases "b1 \<in> supp N1") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1158 | case False | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1159 |         hence "?K = {}" using sM(2) by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1160 | thus ?thesis using False by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1161 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1162 | case True | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1163 | def c \<equiv> "f1 b1" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1164 | have c: "c \<in> supp P" and b1: "b1 \<in> set1 c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1165 | unfolding set1_def c_def P1 using True by (auto simp: mmap_image) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1166 |         have "setsum M ?K = setsum M {a. p1 a = b1 \<and> a \<in> SET}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1167 | apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1168 | also have "... = setsum M ((\<lambda> b2. u c b1 b2) ` (set2 c))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1169 | apply(rule setsum_cong) using c b1 proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1170 | fix a assume p1a: "p1 a \<in> set1 c" and "0 < P c" and "a \<in> SET" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1171 | hence ac: "a \<in> sset c" using p1_rev by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1172 | hence "a = u c (p1 a) (p2 a)" using c by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1173 | moreover have "p2 a \<in> set2 c" using ac c by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1174 | ultimately show "a \<in> u c (p1 a) ` set2 c" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1175 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1176 | fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1177 | hence "u c b1 b2 \<in> SET" using c by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1178 | qed auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1179 | also have "... = setsum (\<lambda> b2. M (u c b1 b2)) (set2 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1180 | unfolding comp_def[symmetric] apply(rule setsum_reindex) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1181 | using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1182 | also have "... = N1 b1" unfolding ss1[OF c b1, symmetric] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1183 | apply(rule setsum_cong[OF refl]) unfolding M_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1184 | using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1] by fastforce | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1185 | finally show ?thesis . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1186 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1187 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1188 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1189 | show "mmap p2 M = N2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1190 | unfolding mmap_def[abs_def] proof(rule ext) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1191 | fix b2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1192 |       let ?K = "{a. p2 a = b2 \<and> 0 < M a}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1193 | show "setsum M ?K = N2 b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1194 | proof(cases "b2 \<in> supp N2") | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1195 | case False | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1196 |         hence "?K = {}" using sM(3) by auto
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1197 | thus ?thesis using False by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1198 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1199 | case True | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1200 | def c \<equiv> "f2 b2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1201 | have c: "c \<in> supp P" and b2: "b2 \<in> set2 c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1202 | unfolding set2_def c_def P2 using True by (auto simp: mmap_image) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1203 |         have "setsum M ?K = setsum M {a. p2 a = b2 \<and> a \<in> SET}"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1204 | apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1205 | also have "... = setsum M ((\<lambda> b1. u c b1 b2) ` (set1 c))" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1206 | apply(rule setsum_cong) using c b2 proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1207 | fix a assume p2a: "p2 a \<in> set2 c" and "0 < P c" and "a \<in> SET" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1208 | hence ac: "a \<in> sset c" using p2_rev by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1209 | hence "a = u c (p1 a) (p2 a)" using c by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1210 | moreover have "p1 a \<in> set1 c" using ac c by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1211 | ultimately show "a \<in> (\<lambda>b1. u c b1 (p2 a)) ` set1 c" by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1212 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1213 | fix b2 assume b1: "b1 \<in> set1 c" and b2: "b2 \<in> set2 c" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1214 | hence "u c b1 b2 \<in> SET" using c by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1215 | qed auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1216 | also have "... = setsum (M o (\<lambda> b1. u c b1 b2)) (set1 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1217 | apply(rule setsum_reindex) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1218 | using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1219 | also have "... = setsum (\<lambda> b1. M (u c b1 b2)) (set1 c)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1220 | unfolding comp_def[symmetric] by simp | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1221 | also have "... = N2 b2" unfolding ss2[OF c b2, symmetric] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1222 | apply(rule setsum_cong[OF refl]) unfolding M_def set2_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1223 | using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2] | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1224 | unfolding set1_def by fastforce | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1225 | finally show ?thesis . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1226 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1227 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1228 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1229 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1230 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1231 | definition multiset_map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1232 | "multiset_map h = Abs_multiset \<circ> mmap h \<circ> count" | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1233 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1234 | bnf_def multiset_map [set_of] "\<lambda>_::'a multiset. natLeq" ["{#}"]
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1235 | unfolding multiset_map_def | 
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1236 | proof - | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1237 | show "Abs_multiset \<circ> mmap id \<circ> count = id" unfolding mmap_id by (auto simp: count_inverse) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1238 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1239 | fix f g | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1240 | show "Abs_multiset \<circ> mmap (g \<circ> f) \<circ> count = | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1241 | Abs_multiset \<circ> mmap g \<circ> count \<circ> (Abs_multiset \<circ> mmap f \<circ> count)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1242 | unfolding comp_def apply(rule ext) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1243 | by (auto simp: Abs_multiset_inverse count mmap_comp1 mmap) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1244 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1245 | fix M f g assume eq: "\<And>a. a \<in> set_of M \<Longrightarrow> f a = g a" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1246 | thus "(Abs_multiset \<circ> mmap f \<circ> count) M = (Abs_multiset \<circ> mmap g \<circ> count) M" apply auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1247 | unfolding cIm_def[abs_def] image_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1248 | by (auto intro!: mmap_cong simp: Abs_multiset_inject count mmap) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1249 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1250 | fix f show "set_of \<circ> (Abs_multiset \<circ> mmap f \<circ> count) = op ` f \<circ> set_of" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1251 | by (auto simp: count mmap mmap_image set_of_Abs_multiset supp_count) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1252 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1253 | show "card_order natLeq" by (rule natLeq_card_order) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1254 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1255 | show "cinfinite natLeq" by (rule natLeq_cinfinite) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1256 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1257 | fix M show "|set_of M| \<le>o natLeq" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1258 | apply(rule ordLess_imp_ordLeq) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1259 | unfolding finite_iff_ordLess_natLeq[symmetric] using finite_set_of . | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1260 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1261 | fix A :: "'a set" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1262 |   have "|{M. set_of M \<subseteq> A}| \<le>o |{as. set as \<subseteq> A}|" using card_of_set_of .
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1263 |   also have "|{as. set as \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1264 | by (rule list_in_bd) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1265 |   finally show "|{M. set_of M \<subseteq> A}| \<le>o ( |A| +c ctwo) ^c natLeq" .
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1266 | next | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1267 | fix A B1 B2 f1 f2 p1 p2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1268 | let ?map = "\<lambda> f. Abs_multiset \<circ> mmap f \<circ> count" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1269 | assume wp: "wpull A B1 B2 f1 f2 p1 p2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1270 |   show "wpull {x. set_of x \<subseteq> A} {x. set_of x \<subseteq> B1} {x. set_of x \<subseteq> B2}
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1271 | (?map f1) (?map f2) (?map p1) (?map p2)" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1272 | unfolding wpull_def proof safe | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1273 | fix y1 y2 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1274 | assume y1: "set_of y1 \<subseteq> B1" and y2: "set_of y2 \<subseteq> B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1275 | and m: "?map f1 y1 = ?map f2 y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1276 | def N1 \<equiv> "count y1" def N2 \<equiv> "count y2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1277 | have "N1 \<in> multiset \<and> supp N1 \<subseteq> B1" and "N2 \<in> multiset \<and> supp N2 \<subseteq> B2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1278 | and "mmap f1 N1 = mmap f2 N2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1279 | using y1 y2 m unfolding N1_def N2_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1280 | by (auto simp: Abs_multiset_inject count mmap) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1281 | then obtain M where M: "M \<in> multiset \<and> supp M \<subseteq> A" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1282 | and N1: "mmap p1 M = N1" and N2: "mmap p2 M = N2" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1283 | using wp_mmap[OF wp] unfolding wpull_def by auto | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1284 | def x \<equiv> "Abs_multiset M" | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1285 |     show "\<exists>x\<in>{x. set_of x \<subseteq> A}. ?map p1 x = y1 \<and> ?map p2 x = y2"
 | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1286 | apply(intro bexI[of _ x]) using M N1 N2 unfolding N1_def N2_def x_def | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1287 | by (auto simp: count_inverse Abs_multiset_inverse) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1288 | qed | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1289 | qed (unfold set_of_empty, auto) | 
| 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1290 | |
| 49514 | 1291 | inductive multiset_rel' where | 
| 1292 | Zero: "multiset_rel' R {#} {#}"
 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1293 | | | 
| 49507 | 1294 | Plus: "\<lbrakk>R a b; multiset_rel' R M N\<rbrakk> \<Longrightarrow> multiset_rel' R (M + {#a#}) (N + {#b#})"
 | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1295 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1296 | lemma multiset_map_Zero_iff[simp]: "multiset_map f M = {#} \<longleftrightarrow> M = {#}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1297 | by (metis image_is_empty multiset.set_natural' set_of_eq_empty_iff) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1298 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1299 | lemma multiset_map_Zero[simp]: "multiset_map f {#} = {#}" by simp
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1300 | |
| 49507 | 1301 | lemma multiset_rel_Zero: "multiset_rel R {#} {#}"
 | 
| 1302 | unfolding multiset_rel_def Gr_def relcomp_unfold by auto | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1303 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1304 | declare multiset.count[simp] | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1305 | declare mmap[simp] | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1306 | declare Abs_multiset_inverse[simp] | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1307 | declare multiset.count_inverse[simp] | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1308 | declare union_preserves_multiset[simp] | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1309 | |
| 49463 | 1310 | lemma mmap_Plus[simp]: | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1311 | assumes "K \<in> multiset" and "L \<in> multiset" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1312 | shows "mmap f (\<lambda>a. K a + L a) a = mmap f K a + mmap f L a" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1313 | proof- | 
| 49463 | 1314 |   have "{aa. f aa = a \<and> (0 < K aa \<or> 0 < L aa)} \<subseteq>
 | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1315 |         {aa. 0 < K aa} \<union> {aa. 0 < L aa}" (is "?C \<subseteq> ?A \<union> ?B") by auto
 | 
| 49463 | 1316 | moreover have "finite (?A \<union> ?B)" apply(rule finite_UnI) | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1317 | using assms unfolding multiset_def by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1318 | ultimately have C: "finite ?C" using finite_subset by blast | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1319 |   have "setsum K {aa. f aa = a \<and> 0 < K aa} = setsum K {aa. f aa = a \<and> 0 < K aa + L aa}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1320 | apply(rule setsum_mono_zero_cong_left) using C by auto | 
| 49463 | 1321 | moreover | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1322 |   have "setsum L {aa. f aa = a \<and> 0 < L aa} = setsum L {aa. f aa = a \<and> 0 < K aa + L aa}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1323 | apply(rule setsum_mono_zero_cong_left) using C by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1324 | ultimately show ?thesis | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1325 | unfolding mmap_def unfolding comm_monoid_add_class.setsum.F_fun_f by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1326 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1327 | |
| 49463 | 1328 | lemma multiset_map_Plus[simp]: | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1329 | "multiset_map f (M1 + M2) = multiset_map f M1 + multiset_map f M2" | 
| 49463 | 1330 | unfolding multiset_map_def | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1331 | apply(subst multiset.count_inject[symmetric]) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1332 | unfolding plus_multiset.rep_eq comp_def by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1333 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1334 | lemma multiset_map_singl[simp]: "multiset_map f {#a#} = {#f a#}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1335 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1336 |   have 0: "\<And> b. card {aa. a = aa \<and> (a = aa \<longrightarrow> f aa = b)} =
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1337 | (if b = f a then 1 else 0)" by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1338 | thus ?thesis | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1339 | unfolding multiset_map_def comp_def mmap_def[abs_def] map_fun_def | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1340 | by (simp, simp add: single_def) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1341 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1342 | |
| 49507 | 1343 | lemma multiset_rel_Plus: | 
| 1344 | assumes ab: "R a b" and MN: "multiset_rel R M N" | |
| 1345 | shows "multiset_rel R (M + {#a#}) (N + {#b#})"
 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1346 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1347 |   {fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
 | 
| 49463 | 1348 |    hence "\<exists>ya. multiset_map fst y + {#a#} = multiset_map fst ya \<and>
 | 
| 1349 |                multiset_map snd y + {#b#} = multiset_map snd ya \<and>
 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1350 |                set_of ya \<subseteq> {(x, y). R x y}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1351 |    apply(intro exI[of _ "y + {#(a,b)#}"]) by auto
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1352 | } | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1353 | thus ?thesis | 
| 49463 | 1354 | using assms | 
| 49507 | 1355 | unfolding multiset_rel_def Gr_def relcomp_unfold by force | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1356 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1357 | |
| 49507 | 1358 | lemma multiset_rel'_imp_multiset_rel: | 
| 1359 | "multiset_rel' R M N \<Longrightarrow> multiset_rel R M N" | |
| 1360 | apply(induct rule: multiset_rel'.induct) | |
| 1361 | using multiset_rel_Zero multiset_rel_Plus by auto | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1362 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1363 | lemma mcard_multiset_map[simp]: "mcard (multiset_map f M) = mcard M" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1364 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1365 |   def A \<equiv> "\<lambda> b. {a. f a = b \<and> a \<in># M}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1366 |   let ?B = "{b. 0 < setsum (count M) (A b)}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1367 |   have "{b. \<exists>a. f a = b \<and> a \<in># M} \<subseteq> f ` {a. a \<in># M}" by auto
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1368 |   moreover have "finite (f ` {a. a \<in># M})" apply(rule finite_imageI)
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1369 | using finite_Collect_mem . | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1370 |   ultimately have fin: "finite {b. \<exists>a. f a = b \<and> a \<in># M}" by(rule finite_subset)
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1371 | have i: "inj_on A ?B" unfolding inj_on_def A_def apply clarsimp | 
| 49463 | 1372 | by (metis (lifting, mono_tags) mem_Collect_eq rel_simps(54) | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1373 | setsum_gt_0_iff setsum_infinite) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1374 | have 0: "\<And> b. 0 < setsum (count M) (A b) \<longleftrightarrow> (\<exists> a \<in> A b. count M a > 0)" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1375 | apply safe | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1376 | apply (metis less_not_refl setsum_gt_0_iff setsum_infinite) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1377 | by (metis A_def finite_Collect_conjI finite_Collect_mem setsum_gt_0_iff) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1378 |   hence AB: "A ` ?B = {A b | b. \<exists> a \<in> A b. count M a > 0}" by auto
 | 
| 49463 | 1379 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1380 | have "setsum (\<lambda> x. setsum (count M) (A x)) ?B = setsum (setsum (count M) o A) ?B" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1381 | unfolding comp_def .. | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1382 | also have "... = (\<Sum>x\<in> A ` ?B. setsum (count M) x)" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1383 | unfolding comm_monoid_add_class.setsum_reindex[OF i, symmetric] .. | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1384 |   also have "... = setsum (count M) (\<Union>x\<in>A ` {b. 0 < setsum (count M) (A b)}. x)"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1385 | (is "_ = setsum (count M) ?J") | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1386 | apply(rule comm_monoid_add_class.setsum_UN_disjoint[symmetric]) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1387 | using 0 fin unfolding A_def by (auto intro!: finite_imageI) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1388 |   also have "?J = {a. a \<in># M}" unfolding AB unfolding A_def by auto
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1389 | finally have "setsum (\<lambda> x. setsum (count M) (A x)) ?B = | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1390 |                 setsum (count M) {a. a \<in># M}" .
 | 
| 49463 | 1391 | thus ?thesis unfolding A_def mcard_def multiset_map_def by (simp add: mmap_def) | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1392 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1393 | |
| 49514 | 1394 | lemma multiset_rel_mcard: | 
| 1395 | assumes "multiset_rel R M N" | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1396 | shows "mcard M = mcard N" | 
| 49507 | 1397 | using assms unfolding multiset_rel_def relcomp_unfold Gr_def by auto | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1398 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1399 | lemma multiset_induct2[case_names empty addL addR]: | 
| 49514 | 1400 | assumes empty: "P {#} {#}"
 | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1401 | and addL: "\<And>M N a. P M N \<Longrightarrow> P (M + {#a#}) N"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1402 | and addR: "\<And>M N a. P M N \<Longrightarrow> P M (N + {#a#})"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1403 | shows "P M N" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1404 | apply(induct N rule: multiset_induct) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1405 | apply(induct M rule: multiset_induct, rule empty, erule addL) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1406 | apply(induct M rule: multiset_induct, erule addR, erule addR) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1407 | done | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1408 | |
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1409 | lemma multiset_induct2_mcard[consumes 1, case_names empty add]: | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1410 | assumes c: "mcard M = mcard N" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1411 | and empty: "P {#} {#}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1412 | and add: "\<And>M N a b. P M N \<Longrightarrow> P (M + {#a#}) (N + {#b#})"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1413 | shows "P M N" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1414 | using c proof(induct M arbitrary: N rule: measure_induct_rule[of mcard]) | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1415 | case (less M) show ?case | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1416 |   proof(cases "M = {#}")
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1417 |     case True hence "N = {#}" using less.prems by auto
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1418 | thus ?thesis using True empty by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1419 | next | 
| 49463 | 1420 |     case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
 | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1421 |     have "N \<noteq> {#}" using False less.prems by auto
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1422 |     then obtain N1 b where N: "N = N1 + {#b#}" by (metis multi_nonempty_split)
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1423 | have "mcard M1 = mcard N1" using less.prems unfolding M N by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1424 | thus ?thesis using M N less.hyps add by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1425 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1426 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1427 | |
| 49463 | 1428 | lemma msed_map_invL: | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1429 | assumes "multiset_map f (M + {#a#}) = N"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1430 | shows "\<exists> N1. N = N1 + {#f a#} \<and> multiset_map f M = N1"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1431 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1432 | have "f a \<in># N" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1433 |   using assms multiset.set_natural'[of f "M + {#a#}"] by auto
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1434 |   then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1435 | have "multiset_map f M = N1" using assms unfolding N by simp | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1436 | thus ?thesis using N by blast | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1437 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1438 | |
| 49463 | 1439 | lemma msed_map_invR: | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1440 | assumes "multiset_map f M = N + {#b#}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1441 | shows "\<exists> M1 a. M = M1 + {#a#} \<and> f a = b \<and> multiset_map f M1 = N"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1442 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1443 | obtain a where a: "a \<in># M" and fa: "f a = b" | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1444 | using multiset.set_natural'[of f M] unfolding assms | 
| 49463 | 1445 | by (metis image_iff mem_set_of_iff union_single_eq_member) | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1446 |   then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1447 | have "multiset_map f M1 = N" using assms unfolding M fa[symmetric] by simp | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1448 | thus ?thesis using M fa by blast | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1449 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1450 | |
| 49507 | 1451 | lemma msed_rel_invL: | 
| 1452 | assumes "multiset_rel R (M + {#a#}) N"
 | |
| 1453 | shows "\<exists> N1 b. N = N1 + {#b#} \<and> R a b \<and> multiset_rel R M N1"
 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1454 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1455 |   obtain K where KM: "multiset_map fst K = M + {#a#}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1456 |   and KN: "multiset_map snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1457 | using assms | 
| 49507 | 1458 | unfolding multiset_rel_def Gr_def relcomp_unfold by auto | 
| 49463 | 1459 |   obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
 | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1460 | and K1M: "multiset_map fst K1 = M" using msed_map_invR[OF KM] by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1461 |   obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "multiset_map snd K1 = N1"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1462 | using msed_map_invL[OF KN[unfolded K]] by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1463 | have Rab: "R a (snd ab)" using sK a unfolding K by auto | 
| 49514 | 1464 | have "multiset_rel R M N1" using sK K1M K1N1 | 
| 49507 | 1465 | unfolding K multiset_rel_def Gr_def relcomp_unfold by auto | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1466 | thus ?thesis using N Rab by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1467 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1468 | |
| 49507 | 1469 | lemma msed_rel_invR: | 
| 1470 | assumes "multiset_rel R M (N + {#b#})"
 | |
| 1471 | shows "\<exists> M1 a. M = M1 + {#a#} \<and> R a b \<and> multiset_rel R M1 N"
 | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1472 | proof- | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1473 |   obtain K where KN: "multiset_map snd K = N + {#b#}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1474 |   and KM: "multiset_map fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1475 | using assms | 
| 49507 | 1476 | unfolding multiset_rel_def Gr_def relcomp_unfold by auto | 
| 49463 | 1477 |   obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
 | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1478 | and K1N: "multiset_map snd K1 = N" using msed_map_invR[OF KN] by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1479 |   obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "multiset_map fst K1 = M1"
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1480 | using msed_map_invL[OF KM[unfolded K]] by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1481 | have Rab: "R (fst ab) b" using sK b unfolding K by auto | 
| 49507 | 1482 | have "multiset_rel R M1 N" using sK K1N K1M1 | 
| 1483 | unfolding K multiset_rel_def Gr_def relcomp_unfold by auto | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1484 | thus ?thesis using M Rab by auto | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1485 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1486 | |
| 49507 | 1487 | lemma multiset_rel_imp_multiset_rel': | 
| 1488 | assumes "multiset_rel R M N" | |
| 1489 | shows "multiset_rel' R M N" | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1490 | using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard]) | 
| 49463 | 1491 | case (less M) | 
| 49507 | 1492 | have c: "mcard M = mcard N" using multiset_rel_mcard[OF less.prems] . | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1493 | show ?case | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1494 |   proof(cases "M = {#}")
 | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1495 |     case True hence "N = {#}" using c by simp
 | 
| 49507 | 1496 | thus ?thesis using True multiset_rel'.Zero by auto | 
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1497 | next | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1498 |     case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
 | 
| 49507 | 1499 |     obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "multiset_rel R M1 N1"
 | 
| 1500 | using msed_rel_invL[OF less.prems[unfolded M]] by auto | |
| 1501 | have "multiset_rel' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp | |
| 1502 | thus ?thesis using multiset_rel'.Plus[of R a b, OF R] unfolding M N by simp | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1503 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1504 | qed | 
| 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1505 | |
| 49507 | 1506 | lemma multiset_rel_multiset_rel': | 
| 1507 | "multiset_rel R M N = multiset_rel' R M N" | |
| 1508 | using multiset_rel_imp_multiset_rel' multiset_rel'_imp_multiset_rel by auto | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1509 | |
| 49507 | 1510 | (* The main end product for multiset_rel: inductive characterization *) | 
| 1511 | theorems multiset_rel_induct[case_names empty add, induct pred: multiset_rel] = | |
| 1512 | multiset_rel'.induct[unfolded multiset_rel_multiset_rel'[symmetric]] | |
| 49440 
4ff2976f4056
Added missing predicators (for multisets and countable sets)
 popescua parents: 
49434diff
changeset | 1513 | |
| 49877 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1514 | |
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1515 | |
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1516 | (* Advanced relator customization *) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1517 | |
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1518 | (* Set vs. sum relators: *) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1519 | (* FIXME: All such facts should be declared as simps: *) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1520 | declare sum_rel_simps[simp] | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1521 | |
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1522 | lemma set_rel_sum_rel[simp]: | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1523 | "set_rel (sum_rel \<chi> \<phi>) A1 A2 \<longleftrightarrow> | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1524 | set_rel \<chi> (Inl -` A1) (Inl -` A2) \<and> set_rel \<phi> (Inr -` A1) (Inr -` A2)" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1525 | (is "?L \<longleftrightarrow> ?Rl \<and> ?Rr") | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1526 | proof safe | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1527 | assume L: "?L" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1528 | show ?Rl unfolding set_rel_def Bex_def vimage_eq proof safe | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1529 | fix l1 assume "Inl l1 \<in> A1" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1530 | then obtain a2 where a2: "a2 \<in> A2" and "sum_rel \<chi> \<phi> (Inl l1) a2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1531 | using L unfolding set_rel_def by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1532 | then obtain l2 where "a2 = Inl l2 \<and> \<chi> l1 l2" by (cases a2, auto) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1533 | thus "\<exists> l2. Inl l2 \<in> A2 \<and> \<chi> l1 l2" using a2 by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1534 | next | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1535 | fix l2 assume "Inl l2 \<in> A2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1536 | then obtain a1 where a1: "a1 \<in> A1" and "sum_rel \<chi> \<phi> a1 (Inl l2)" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1537 | using L unfolding set_rel_def by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1538 | then obtain l1 where "a1 = Inl l1 \<and> \<chi> l1 l2" by (cases a1, auto) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1539 | thus "\<exists> l1. Inl l1 \<in> A1 \<and> \<chi> l1 l2" using a1 by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1540 | qed | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1541 | show ?Rr unfolding set_rel_def Bex_def vimage_eq proof safe | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1542 | fix r1 assume "Inr r1 \<in> A1" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1543 | then obtain a2 where a2: "a2 \<in> A2" and "sum_rel \<chi> \<phi> (Inr r1) a2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1544 | using L unfolding set_rel_def by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1545 | then obtain r2 where "a2 = Inr r2 \<and> \<phi> r1 r2" by (cases a2, auto) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1546 | thus "\<exists> r2. Inr r2 \<in> A2 \<and> \<phi> r1 r2" using a2 by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1547 | next | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1548 | fix r2 assume "Inr r2 \<in> A2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1549 | then obtain a1 where a1: "a1 \<in> A1" and "sum_rel \<chi> \<phi> a1 (Inr r2)" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1550 | using L unfolding set_rel_def by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1551 | then obtain r1 where "a1 = Inr r1 \<and> \<phi> r1 r2" by (cases a1, auto) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1552 | thus "\<exists> r1. Inr r1 \<in> A1 \<and> \<phi> r1 r2" using a1 by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1553 | qed | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1554 | next | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1555 | assume Rl: "?Rl" and Rr: "?Rr" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1556 | show ?L unfolding set_rel_def Bex_def vimage_eq proof safe | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1557 | fix a1 assume a1: "a1 \<in> A1" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1558 | show "\<exists> a2. a2 \<in> A2 \<and> sum_rel \<chi> \<phi> a1 a2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1559 | proof(cases a1) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1560 | case (Inl l1) then obtain l2 where "Inl l2 \<in> A2 \<and> \<chi> l1 l2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1561 | using Rl a1 unfolding set_rel_def by blast | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1562 | thus ?thesis unfolding Inl by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1563 | next | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1564 | case (Inr r1) then obtain r2 where "Inr r2 \<in> A2 \<and> \<phi> r1 r2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1565 | using Rr a1 unfolding set_rel_def by blast | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1566 | thus ?thesis unfolding Inr by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1567 | qed | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1568 | next | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1569 | fix a2 assume a2: "a2 \<in> A2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1570 | show "\<exists> a1. a1 \<in> A1 \<and> sum_rel \<chi> \<phi> a1 a2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1571 | proof(cases a2) | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1572 | case (Inl l2) then obtain l1 where "Inl l1 \<in> A1 \<and> \<chi> l1 l2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1573 | using Rl a2 unfolding set_rel_def by blast | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1574 | thus ?thesis unfolding Inl by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1575 | next | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1576 | case (Inr r2) then obtain r1 where "Inr r1 \<in> A1 \<and> \<phi> r1 r2" | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1577 | using Rr a2 unfolding set_rel_def by blast | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1578 | thus ?thesis unfolding Inr by auto | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1579 | qed | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1580 | qed | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1581 | qed | 
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1582 | |
| 
b75555ec30a4
ported HOL/BNF/Examples/Derivation_Trees to the latest status of the codatatype package
 popescua parents: 
49514diff
changeset | 1583 | |
| 49309 
f20b24214ac2
split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
 blanchet parents: diff
changeset | 1584 | end |