# HG changeset patch # User blanchet # Date 1347420107 -7200 # Node ID f20b24214ac2f725e535ec9c3db06aee4a1d43db # Parent 6190b701e4f49c4867b9d2cde0a5ba46927436bf split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/BNF_Comp.thy --- a/src/HOL/Codatatype/BNF_Comp.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/BNF_Comp.thy Wed Sep 12 05:21:47 2012 +0200 @@ -9,9 +9,9 @@ theory BNF_Comp imports Basic_BNFs -uses - "Tools/bnf_comp_tactics.ML" - "Tools/bnf_comp.ML" begin +ML_file "Tools/bnf_comp_tactics.ML" +ML_file "Tools/bnf_comp.ML" + end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/BNF_Def.thy --- a/src/HOL/Codatatype/BNF_Def.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/BNF_Def.thy Wed Sep 12 05:21:47 2012 +0200 @@ -12,9 +12,9 @@ keywords "print_bnfs" :: diag and "bnf_def" :: thy_goal -uses - "Tools/bnf_def_tactics.ML" - "Tools/bnf_def.ML" begin +ML_file "Tools/bnf_def_tactics.ML" +ML_file"Tools/bnf_def.ML" + end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/BNF_FP.thy --- a/src/HOL/Codatatype/BNF_FP.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/BNF_FP.thy Wed Sep 12 05:21:47 2012 +0200 @@ -12,10 +12,10 @@ imports BNF_Comp BNF_Wrap keywords "defaults" -uses - "Tools/bnf_fp_util.ML" - "Tools/bnf_fp_sugar_tactics.ML" - "Tools/bnf_fp_sugar.ML" begin +ML_file "Tools/bnf_fp_util.ML" +ML_file "Tools/bnf_fp_sugar_tactics.ML" +ML_file "Tools/bnf_fp_sugar.ML" + end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/BNF_GFP.thy --- a/src/HOL/Codatatype/BNF_GFP.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/BNF_GFP.thy Wed Sep 12 05:21:47 2012 +0200 @@ -12,10 +12,10 @@ keywords "codata_raw" :: thy_decl and "codata" :: thy_decl -uses - "Tools/bnf_gfp_util.ML" - "Tools/bnf_gfp_tactics.ML" - "Tools/bnf_gfp.ML" begin +ML_file "Tools/bnf_gfp_util.ML" +ML_file "Tools/bnf_gfp_tactics.ML" +ML_file "Tools/bnf_gfp.ML" + end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/BNF_LFP.thy --- a/src/HOL/Codatatype/BNF_LFP.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/BNF_LFP.thy Wed Sep 12 05:21:47 2012 +0200 @@ -12,10 +12,10 @@ keywords "data_raw" :: thy_decl and "data" :: thy_decl -uses - "Tools/bnf_lfp_util.ML" - "Tools/bnf_lfp_tactics.ML" - "Tools/bnf_lfp.ML" begin +ML_file "Tools/bnf_lfp_util.ML" +ML_file "Tools/bnf_lfp_tactics.ML" +ML_file "Tools/bnf_lfp.ML" + end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/BNF_Util.thy --- a/src/HOL/Codatatype/BNF_Util.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/BNF_Util.thy Wed Sep 12 05:21:47 2012 +0200 @@ -13,8 +13,6 @@ "../Ordinals_and_Cardinals/Cardinal_Arithmetic" "~~/src/HOL/Library/Prefix_Order" Equiv_Relations_More -uses - ("Tools/bnf_util.ML") begin lemma iffI_np: "\x \ \ y; \ x \ y\ \ \ x \ y" diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/BNF_Wrap.thy --- a/src/HOL/Codatatype/BNF_Wrap.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/BNF_Wrap.thy Wed Sep 12 05:21:47 2012 +0200 @@ -12,9 +12,9 @@ keywords "wrap_data" :: thy_goal and "no_dests" -uses - "Tools/bnf_wrap_tactics.ML" - "Tools/bnf_wrap.ML" begin +ML_file "Tools/bnf_wrap_tactics.ML" +ML_file "Tools/bnf_wrap.ML" + end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/Basic_BNFs.thy --- a/src/HOL/Codatatype/Basic_BNFs.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/Basic_BNFs.thy Wed Sep 12 05:21:47 2012 +0200 @@ -4,10 +4,10 @@ Author: Jasmin Blanchette, TU Muenchen Copyright 2012 -Registration of various types as bounded natural functors. +Registration of basic types as bounded natural functors. *) -header {* Registration of Various Types as Bounded Natural Functors *} +header {* Registration of Basic Types as Bounded Natural Functors *} theory Basic_BNFs imports BNF_Def "~~/src/HOL/Quotient_Examples/FSet" "~~/src/HOL/Library/Multiset" Countable_Set @@ -71,7 +71,6 @@ unfolding DEADID_pred_def DEADID.rel_Id by simp ML {* - signature BASIC_BNFS = sig val ID_bnf: BNF_Def.BNF @@ -84,13 +83,12 @@ structure Basic_BNFs : BASIC_BNFS = struct - val ID_bnf = the (BNF_Def.bnf_of @{context} "ID"); - val DEADID_bnf = the (BNF_Def.bnf_of @{context} "DEADID"); +val ID_bnf = the (BNF_Def.bnf_of @{context} "ID"); +val DEADID_bnf = the (BNF_Def.bnf_of @{context} "DEADID"); - val rel_def = BNF_Def.rel_def_of_bnf ID_bnf; - val ID_rel_def = rel_def RS sym; - val ID_pred_def = - Local_Defs.unfold @{context} [rel_def] (BNF_Def.pred_def_of_bnf ID_bnf) RS sym; +val rel_def = BNF_Def.rel_def_of_bnf ID_bnf; +val ID_rel_def = rel_def RS sym; +val ID_pred_def = Local_Defs.unfold @{context} [rel_def] (BNF_Def.pred_def_of_bnf ID_bnf) RS sym; end; *} @@ -422,1103 +420,4 @@ unfolding fun_rel_def fun_pred_def Gr_def relcomp_unfold converse_unfold by (auto intro!: exI dest!: in_mono) -lemma card_of_list_in: - "|{xs. set xs \ A}| \o |Pfunc (UNIV :: nat set) A|" (is "|?LHS| \o |?RHS|") -proof - - let ?f = "%xs. %i. if i < length xs \ set xs \ A then Some (nth xs i) else None" - have "inj_on ?f ?LHS" unfolding inj_on_def fun_eq_iff - proof safe - fix xs :: "'a list" and ys :: "'a list" - assume su: "set xs \ A" "set ys \ A" and eq: "\i. ?f xs i = ?f ys i" - hence *: "length xs = length ys" - by (metis linorder_cases option.simps(2) order_less_irrefl) - thus "xs = ys" by (rule nth_equalityI) (metis * eq su option.inject) - qed - moreover have "?f ` ?LHS \ ?RHS" unfolding Pfunc_def by fastforce - ultimately show ?thesis using card_of_ordLeq by blast -qed - -lemma list_in_empty: "A = {} \ {x. set x \ A} = {[]}" -by simp - -lemma card_of_Func: "|Func A B| =o |B| ^c |A|" -unfolding cexp_def Field_card_of by (rule card_of_refl) - -lemma not_emp_czero_notIn_ordIso_Card_order: -"A \ {} \ ( |A|, czero) \ ordIso \ Card_order |A|" - apply (rule conjI) - apply (metis Field_card_of czeroE) - by (rule card_of_Card_order) - -lemma list_in_bd: "|{x. set x \ A}| \o ( |A| +c ctwo) ^c natLeq" -proof - - fix A :: "'a set" - show "|{x. set x \ A}| \o ( |A| +c ctwo) ^c natLeq" - proof (cases "A = {}") - case False thus ?thesis - apply - - apply (rule ordLeq_transitive) - apply (rule card_of_list_in) - apply (rule ordLeq_transitive) - apply (erule card_of_Pfunc_Pow_Func) - apply (rule ordIso_ordLeq_trans) - apply (rule Times_cprod) - apply (rule cprod_cinfinite_bound) - apply (rule ordIso_ordLeq_trans) - apply (rule Pow_cexp_ctwo) - apply (rule ordIso_ordLeq_trans) - apply (rule cexp_cong2) - apply (rule card_of_nat) - apply (rule Card_order_ctwo) - apply (rule card_of_Card_order) - apply (rule natLeq_Card_order) - apply (rule disjI1) - apply (rule ctwo_Cnotzero) - apply (rule cexp_mono1) - apply (rule ordLeq_csum2) - apply (rule Card_order_ctwo) - apply (rule disjI1) - apply (rule ctwo_Cnotzero) - apply (rule natLeq_Card_order) - apply (rule ordIso_ordLeq_trans) - apply (rule card_of_Func) - apply (rule ordIso_ordLeq_trans) - apply (rule cexp_cong2) - apply (rule card_of_nat) - apply (rule card_of_Card_order) - apply (rule card_of_Card_order) - apply (rule natLeq_Card_order) - apply (rule disjI1) - apply (erule not_emp_czero_notIn_ordIso_Card_order) - apply (rule cexp_mono1) - apply (rule ordLeq_csum1) - apply (rule card_of_Card_order) - apply (rule disjI1) - apply (erule not_emp_czero_notIn_ordIso_Card_order) - apply (rule natLeq_Card_order) - apply (rule card_of_Card_order) - apply (rule card_of_Card_order) - apply (rule Cinfinite_cexp) - apply (rule ordLeq_csum2) - apply (rule Card_order_ctwo) - apply (rule conjI) - apply (rule natLeq_cinfinite) - by (rule natLeq_Card_order) - next - case True thus ?thesis - apply - - apply (rule ordIso_ordLeq_trans) - apply (rule card_of_ordIso_subst) - apply (erule list_in_empty) - apply (rule ordIso_ordLeq_trans) - apply (rule single_cone) - apply (rule cone_ordLeq_cexp) - apply (rule ordLeq_transitive) - apply (rule cone_ordLeq_ctwo) - apply (rule ordLeq_csum2) - by (rule Card_order_ctwo) - qed -qed - -bnf_def list = map [set] "\_::'a list. natLeq" ["[]"] -proof - - show "map id = id" by (rule List.map.id) -next - fix f g - show "map (g o f) = map g o map f" by (rule List.map.comp[symmetric]) -next - fix x f g - assume "\z. z \ set x \ f z = g z" - thus "map f x = map g x" by simp -next - fix f - show "set o map f = image f o set" by (rule ext, unfold o_apply, rule set_map) -next - show "card_order natLeq" by (rule natLeq_card_order) -next - show "cinfinite natLeq" by (rule natLeq_cinfinite) -next - fix x - show "|set x| \o natLeq" - apply (rule ordLess_imp_ordLeq) - apply (rule finite_ordLess_infinite[OF _ natLeq_Well_order]) - unfolding Field_natLeq Field_card_of by (auto simp: card_of_well_order_on) -next - fix A :: "'a set" - show "|{x. set x \ A}| \o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd) -next - fix A B1 B2 f1 f2 p1 p2 - assume "wpull A B1 B2 f1 f2 p1 p2" - hence pull: "\b1 b2. b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2 \ \a \ A. p1 a = b1 \ p2 a = b2" - unfolding wpull_def by auto - show "wpull {x. set x \ A} {x. set x \ B1} {x. set x \ B2} (map f1) (map f2) (map p1) (map p2)" - (is "wpull ?A ?B1 ?B2 _ _ _ _") - proof (unfold wpull_def) - { fix as bs assume *: "as \ ?B1" "bs \ ?B2" "map f1 as = map f2 bs" - hence "length as = length bs" by (metis length_map) - hence "\zs \ ?A. map p1 zs = as \ map p2 zs = bs" using * - proof (induct as bs rule: list_induct2) - case (Cons a as b bs) - hence "a \ B1" "b \ B2" "f1 a = f2 b" by auto - with pull obtain z where "z \ A" "p1 z = a" "p2 z = b" by blast - moreover - from Cons obtain zs where "zs \ ?A" "map p1 zs = as" "map p2 zs = bs" by auto - ultimately have "z # zs \ ?A" "map p1 (z # zs) = a # as \ map p2 (z # zs) = b # bs" by auto - thus ?case by (rule_tac x = "z # zs" in bexI) - qed simp - } - thus "\as bs. as \ ?B1 \ bs \ ?B2 \ map f1 as = map f2 bs \ - (\zs \ ?A. map p1 zs = as \ map p2 zs = bs)" by blast - qed -qed auto - -(* Finite sets *) -abbreviation afset where "afset \ abs_fset" -abbreviation rfset where "rfset \ rep_fset" - -lemma fset_fset_member: -"fset A = {a. a |\| A}" -unfolding fset_def fset_member_def by auto - -lemma afset_rfset: -"afset (rfset x) = x" -by (rule Quotient_fset[unfolded Quotient_def, THEN conjunct1, rule_format]) - -lemma afset_rfset_id: -"afset o rfset = id" -unfolding comp_def afset_rfset id_def .. - -lemma rfset: -"rfset A = rfset B \ A = B" -by (metis afset_rfset) - -lemma afset_set: -"afset as = afset bs \ set as = set bs" -using Quotient_fset unfolding Quotient_def list_eq_def by auto - -lemma surj_afset: -"\ as. A = afset as" -by (metis afset_rfset) - -lemma fset_def2: -"fset = set o rfset" -unfolding fset_def map_fun_def[abs_def] by simp - -lemma fset_def2_raw: -"fset A = set (rfset A)" -unfolding fset_def2 by simp - -lemma fset_comp_afset: -"fset o afset = set" -unfolding fset_def2 comp_def apply(rule ext) -unfolding afset_set[symmetric] afset_rfset .. - -lemma fset_afset: -"fset (afset as) = set as" -unfolding fset_comp_afset[symmetric] by simp - -lemma set_rfset_afset: -"set (rfset (afset as)) = set as" -unfolding afset_set[symmetric] afset_rfset .. - -lemma map_fset_comp_afset: -"(map_fset f) o afset = afset o (map f)" -unfolding map_fset_def map_fun_def[abs_def] comp_def apply(rule ext) -unfolding afset_set set_map set_rfset_afset id_apply .. - -lemma map_fset_afset: -"(map_fset f) (afset as) = afset (map f as)" -using map_fset_comp_afset unfolding comp_def fun_eq_iff by auto - -lemma fset_map_fset: -"fset (map_fset f A) = (image f) (fset A)" -apply(subst afset_rfset[symmetric, of A]) -unfolding map_fset_afset fset_afset set_map -unfolding fset_def2_raw .. - -lemma map_fset_def2: -"map_fset f = afset o (map f) o rfset" -unfolding map_fset_def map_fun_def[abs_def] by simp - -lemma map_fset_def2_raw: -"map_fset f A = afset (map f (rfset A))" -unfolding map_fset_def2 by simp - -lemma finite_ex_fset: -assumes "finite A" -shows "\ B. fset B = A" -by (metis assms finite_list fset_afset) - -lemma wpull_image: -assumes "wpull A B1 B2 f1 f2 p1 p2" -shows "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)" -unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify - fix Y1 Y2 assume Y1: "Y1 \ B1" and Y2: "Y2 \ B2" and EQ: "f1 ` Y1 = f2 ` Y2" - def X \ "{a \ A. p1 a \ Y1 \ p2 a \ Y2}" - show "\X\A. p1 ` X = Y1 \ p2 ` X = Y2" - proof (rule exI[of _ X], intro conjI) - show "p1 ` X = Y1" - proof - show "Y1 \ p1 ` X" - proof safe - fix y1 assume y1: "y1 \ Y1" - then obtain y2 where y2: "y2 \ Y2" and eq: "f1 y1 = f2 y2" using EQ by auto - then obtain x where "x \ A" and "p1 x = y1" and "p2 x = y2" - using assms y1 Y1 Y2 unfolding wpull_def by blast - thus "y1 \ p1 ` X" unfolding X_def using y1 y2 by auto - qed - qed(unfold X_def, auto) - show "p2 ` X = Y2" - proof - show "Y2 \ p2 ` X" - proof safe - fix y2 assume y2: "y2 \ Y2" - then obtain y1 where y1: "y1 \ Y1" and eq: "f1 y1 = f2 y2" using EQ by force - then obtain x where "x \ A" and "p1 x = y1" and "p2 x = y2" - using assms y2 Y1 Y2 unfolding wpull_def by blast - thus "y2 \ p2 ` X" unfolding X_def using y1 y2 by auto - qed - qed(unfold X_def, auto) - qed(unfold X_def, auto) -qed - -lemma fset_to_fset: "finite A \ fset (the_inv fset A) = A" -by (rule f_the_inv_into_f) (auto simp: inj_on_def fset_cong dest!: finite_ex_fset) - -bnf_def fset = map_fset [fset] "\_::'a fset. natLeq" ["{||}"] -proof - - show "map_fset id = id" - unfolding map_fset_def2 map_id o_id afset_rfset_id .. -next - fix f g - show "map_fset (g o f) = map_fset g o map_fset f" - unfolding map_fset_def2 map.comp[symmetric] comp_def apply(rule ext) - unfolding afset_set set_map fset_def2_raw[symmetric] image_image[symmetric] - unfolding map_fset_afset[symmetric] map_fset_image afset_rfset - by (rule refl) -next - fix x f g - assume "\z. z \ fset x \ f z = g z" - hence "map f (rfset x) = map g (rfset x)" - apply(intro map_cong) unfolding fset_def2_raw by auto - thus "map_fset f x = map_fset g x" unfolding map_fset_def2_raw - by (rule arg_cong) -next - fix f - show "fset o map_fset f = image f o fset" - unfolding comp_def fset_map_fset .. -next - show "card_order natLeq" by (rule natLeq_card_order) -next - show "cinfinite natLeq" by (rule natLeq_cinfinite) -next - fix x - show "|fset x| \o natLeq" - unfolding fset_def2_raw - apply (rule ordLess_imp_ordLeq) - apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) - by (rule finite_set) -next - fix A :: "'a set" - have "|{x. fset x \ A}| \o |afset ` {as. set as \ A}|" - apply(rule card_of_mono1) unfolding fset_def2_raw apply auto - apply (rule image_eqI) - by (auto simp: afset_rfset) - also have "|afset ` {as. set as \ A}| \o |{as. set as \ A}|" using card_of_image . - also have "|{as. set as \ A}| \o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd) - finally show "|{x. fset x \ A}| \o ( |A| +c ctwo) ^c natLeq" . -next - fix A B1 B2 f1 f2 p1 p2 - assume wp: "wpull A B1 B2 f1 f2 p1 p2" - hence "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)" - by(rule wpull_image) - show "wpull {x. fset x \ A} {x. fset x \ B1} {x. fset x \ B2} - (map_fset f1) (map_fset f2) (map_fset p1) (map_fset p2)" - unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify - fix y1 y2 - assume Y1: "fset y1 \ B1" and Y2: "fset y2 \ B2" - assume "map_fset f1 y1 = map_fset f2 y2" - hence EQ: "f1 ` (fset y1) = f2 ` (fset y2)" unfolding map_fset_def2_raw - unfolding afset_set set_map fset_def2_raw . - with Y1 Y2 obtain X where X: "X \ A" - and Y1: "p1 ` X = fset y1" and Y2: "p2 ` X = fset y2" - using wpull_image[OF wp] unfolding wpull_def Pow_def - unfolding Bex_def mem_Collect_eq apply - - apply(erule allE[of _ "fset y1"], erule allE[of _ "fset y2"]) by auto - have "\ y1' \ fset y1. \ x. x \ X \ y1' = p1 x" using Y1 by auto - then obtain q1 where q1: "\ y1' \ fset y1. q1 y1' \ X \ y1' = p1 (q1 y1')" by metis - have "\ y2' \ fset y2. \ x. x \ X \ y2' = p2 x" using Y2 by auto - then obtain q2 where q2: "\ y2' \ fset y2. q2 y2' \ X \ y2' = p2 (q2 y2')" by metis - def X' \ "q1 ` (fset y1) \ q2 ` (fset y2)" - have X': "X' \ A" and Y1: "p1 ` X' = fset y1" and Y2: "p2 ` X' = fset y2" - using X Y1 Y2 q1 q2 unfolding X'_def by auto - have fX': "finite X'" unfolding X'_def by simp - then obtain x where X'eq: "X' = fset x" by (auto dest: finite_ex_fset) - show "\x. fset x \ A \ map_fset p1 x = y1 \ map_fset p2 x = y2" - apply(intro exI[of _ "x"]) using X' Y1 Y2 - unfolding X'eq map_fset_def2_raw fset_def2_raw set_map[symmetric] - afset_set[symmetric] afset_rfset by simp - qed -qed auto - -lemma fset_pred[simp]: "fset_pred R a b \ - ((\t \ fset a. (\u \ fset b. R t u)) \ - (\t \ fset b. (\u \ fset a. R u t)))" (is "?L = ?R") -proof - assume ?L thus ?R unfolding fset_rel_def fset_pred_def - Gr_def relcomp_unfold converse_unfold - apply (simp add: subset_eq Ball_def) - apply (rule conjI) - apply (clarsimp, metis snd_conv) - by (clarsimp, metis fst_conv) -next - assume ?R - def R' \ "the_inv fset (Collect (split R) \ (fset a \ fset b))" (is "the_inv fset ?R'") - have "finite ?R'" by (intro finite_Int[OF disjI2] finite_cartesian_product) auto - hence *: "fset R' = ?R'" unfolding R'_def by (intro fset_to_fset) - show ?L unfolding fset_rel_def fset_pred_def Gr_def relcomp_unfold converse_unfold - proof (intro CollectI prod_caseI exI conjI) - from * show "(R', a) = (R', map_fset fst R')" using conjunct1[OF `?R`] - by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits) - from * show "(R', b) = (R', map_fset snd R')" using conjunct2[OF `?R`] - by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits) - qed (auto simp add: *) -qed - -(* Countable sets *) - -lemma card_of_countable_sets_range: -fixes A :: "'a set" -shows "|{X. X \ A \ countable X \ X \ {}}| \o |{f::nat \ 'a. range f \ A}|" -apply(rule card_of_ordLeqI[of fromNat]) using inj_on_fromNat -unfolding inj_on_def by auto - -lemma card_of_countable_sets_Func: -"|{X. X \ A \ countable X \ X \ {}}| \o |A| ^c natLeq" -using card_of_countable_sets_range card_of_Func_UNIV[THEN ordIso_symmetric] -unfolding cexp_def Field_natLeq Field_card_of -by(rule ordLeq_ordIso_trans) - -lemma ordLeq_countable_subsets: -"|A| \o |{X. X \ A \ countable X}|" -apply(rule card_of_ordLeqI[of "\ a. {a}"]) unfolding inj_on_def by auto - -lemma finite_countable_subset: -"finite {X. X \ A \ countable X} \ finite A" -apply default - apply (erule contrapos_pp) - apply (rule card_of_ordLeq_infinite) - apply (rule ordLeq_countable_subsets) - apply assumption -apply (rule finite_Collect_conjI) -apply (rule disjI1) -by (erule finite_Collect_subsets) - -lemma card_of_countable_sets: -"|{X. X \ A \ countable X}| \o ( |A| +c ctwo) ^c natLeq" -(is "|?L| \o _") -proof(cases "finite A") - let ?R = "Func (UNIV::nat set) (A <+> (UNIV::bool set))" - case True hence "finite ?L" by simp - moreover have "infinite ?R" - apply(rule infinite_Func[of _ "Inr True" "Inr False"]) by auto - ultimately show ?thesis unfolding cexp_def csum_def ctwo_def Field_natLeq Field_card_of - apply(intro ordLess_imp_ordLeq) by (rule finite_ordLess_infinite2) -next - case False - hence "|{X. X \ A \ countable X}| =o |{X. X \ A \ countable X} - {{}}|" - by (intro card_of_infinite_diff_finitte finite.emptyI finite.insertI ordIso_symmetric) - (unfold finite_countable_subset) - also have "|{X. X \ A \ countable X} - {{}}| \o |A| ^c natLeq" - using card_of_countable_sets_Func[of A] unfolding set_diff_eq by auto - also have "|A| ^c natLeq \o ( |A| +c ctwo) ^c natLeq" - apply(rule cexp_mono1_cone_ordLeq) - apply(rule ordLeq_csum1, rule card_of_Card_order) - apply (rule cone_ordLeq_cexp) - apply (rule cone_ordLeq_Cnotzero) - using csum_Cnotzero2 ctwo_Cnotzero apply blast - by (rule natLeq_Card_order) - finally show ?thesis . -qed - -bnf_def cset = cIm [rcset] "\_::'a cset. natLeq" ["cEmp"] -proof - - show "cIm id = id" unfolding cIm_def[abs_def] id_def by auto -next - fix f g show "cIm (g \ f) = cIm g \ cIm f" - unfolding cIm_def[abs_def] apply(rule ext) unfolding comp_def by auto -next - fix C f g assume eq: "\a. a \ rcset C \ f a = g a" - thus "cIm f C = cIm g C" - unfolding cIm_def[abs_def] unfolding image_def by auto -next - fix f show "rcset \ cIm f = op ` f \ rcset" unfolding cIm_def[abs_def] by auto -next - show "card_order natLeq" by (rule natLeq_card_order) -next - show "cinfinite natLeq" by (rule natLeq_cinfinite) -next - fix C show "|rcset C| \o natLeq" using rcset unfolding countable_def . -next - fix A :: "'a set" - have "|{Z. rcset Z \ A}| \o |acset ` {X. X \ A \ countable X}|" - apply(rule card_of_mono1) unfolding Pow_def image_def - proof (rule Collect_mono, clarsimp) - fix x - assume "rcset x \ A" - hence "rcset x \ A \ countable (rcset x) \ x = acset (rcset x)" - using acset_rcset[of x] rcset[of x] by force - thus "\y \ A. countable y \ x = acset y" by blast - qed - also have "|acset ` {X. X \ A \ countable X}| \o |{X. X \ A \ countable X}|" - using card_of_image . - also have "|{X. X \ A \ countable X}| \o ( |A| +c ctwo) ^c natLeq" - using card_of_countable_sets . - finally show "|{Z. rcset Z \ A}| \o ( |A| +c ctwo) ^c natLeq" . -next - fix A B1 B2 f1 f2 p1 p2 - assume wp: "wpull A B1 B2 f1 f2 p1 p2" - show "wpull {x. rcset x \ A} {x. rcset x \ B1} {x. rcset x \ B2} - (cIm f1) (cIm f2) (cIm p1) (cIm p2)" - unfolding wpull_def proof safe - fix y1 y2 - assume Y1: "rcset y1 \ B1" and Y2: "rcset y2 \ B2" - assume "cIm f1 y1 = cIm f2 y2" - hence EQ: "f1 ` (rcset y1) = f2 ` (rcset y2)" - unfolding cIm_def by auto - with Y1 Y2 obtain X where X: "X \ A" - and Y1: "p1 ` X = rcset y1" and Y2: "p2 ` X = rcset y2" - using wpull_image[OF wp] unfolding wpull_def Pow_def - unfolding Bex_def mem_Collect_eq apply - - apply(erule allE[of _ "rcset y1"], erule allE[of _ "rcset y2"]) by auto - have "\ y1' \ rcset y1. \ x. x \ X \ y1' = p1 x" using Y1 by auto - then obtain q1 where q1: "\ y1' \ rcset y1. q1 y1' \ X \ y1' = p1 (q1 y1')" by metis - have "\ y2' \ rcset y2. \ x. x \ X \ y2' = p2 x" using Y2 by auto - then obtain q2 where q2: "\ y2' \ rcset y2. q2 y2' \ X \ y2' = p2 (q2 y2')" by metis - def X' \ "q1 ` (rcset y1) \ q2 ` (rcset y2)" - have X': "X' \ A" and Y1: "p1 ` X' = rcset y1" and Y2: "p2 ` X' = rcset y2" - using X Y1 Y2 q1 q2 unfolding X'_def by fast+ - have fX': "countable X'" unfolding X'_def by simp - then obtain x where X'eq: "X' = rcset x" by (metis rcset_acset) - show "\x\{x. rcset x \ A}. cIm p1 x = y1 \ cIm p2 x = y2" - apply(intro bexI[of _ "x"]) using X' Y1 Y2 unfolding X'eq cIm_def by auto - qed -qed (unfold cEmp_def, auto) - - -(* Multisets *) - -lemma setsum_gt_0_iff: -fixes f :: "'a \ nat" assumes "finite A" -shows "setsum f A > 0 \ (\ a \ A. f a > 0)" -(is "?L \ ?R") -proof- - have "?L \ \ setsum f A = 0" by fast - also have "... \ (\ a \ A. f a \ 0)" using assms by simp - also have "... \ ?R" by simp - finally show ?thesis . -qed - -(* *) -definition mmap :: "('a \ 'b) \ ('a \ nat) \ 'b \ nat" where -"mmap h f b = setsum f {a. h a = b \ f a > 0}" - -lemma mmap_id: "mmap id = id" -proof (rule ext)+ - fix f a show "mmap id f a = id f a" - proof(cases "f a = 0") - case False - hence 1: "{aa. aa = a \ 0 < f aa} = {a}" by auto - show ?thesis by (simp add: mmap_def id_apply 1) - qed(unfold mmap_def, auto) -qed - -lemma inj_on_setsum_inv: -assumes f: "f \ multiset" -and 1: "(0::nat) < setsum f {a. h a = b' \ 0 < f a}" (is "0 < setsum f ?A'") -and 2: "{a. h a = b \ 0 < f a} = {a. h a = b' \ 0 < f a}" (is "?A = ?A'") -shows "b = b'" -proof- - have "finite ?A'" using f unfolding multiset_def by auto - hence "?A' \ {}" using 1 setsum_gt_0_iff by auto - thus ?thesis using 2 by auto -qed - -lemma mmap_comp: -fixes h1 :: "'a \ 'b" and h2 :: "'b \ 'c" -assumes f: "f \ multiset" -shows "mmap (h2 o h1) f = (mmap h2 o mmap h1) f" -unfolding mmap_def[abs_def] comp_def proof(rule ext)+ - fix c :: 'c - let ?A = "{a. h2 (h1 a) = c \ 0 < f a}" - let ?As = "\ b. {a. h1 a = b \ 0 < f a}" - let ?B = "{b. h2 b = c \ 0 < setsum f (?As b)}" - have 0: "{?As b | b. b \ ?B} = ?As ` ?B" by auto - have "\ b. finite (?As b)" using f unfolding multiset_def by simp - hence "?B = {b. h2 b = c \ ?As b \ {}}" using setsum_gt_0_iff by auto - hence A: "?A = \ {?As b | b. b \ ?B}" by auto - have "setsum f ?A = setsum (setsum f) {?As b | b. b \ ?B}" - unfolding A apply(rule setsum_Union_disjoint) - using f unfolding multiset_def by auto - also have "... = setsum (setsum f) (?As ` ?B)" unfolding 0 .. - also have "... = setsum (setsum f o ?As) ?B" apply(rule setsum_reindex) - unfolding inj_on_def apply auto using inj_on_setsum_inv[OF f, of h1] by blast - also have "... = setsum (\ b. setsum f (?As b)) ?B" unfolding comp_def .. - finally show "setsum f ?A = setsum (\ b. setsum f (?As b)) ?B" . -qed - -lemma mmap_comp1: -fixes h1 :: "'a \ 'b" and h2 :: "'b \ 'c" -assumes "f \ multiset" -shows "mmap (\ a. h2 (h1 a)) f = mmap h2 (mmap h1 f)" -using mmap_comp[OF assms] unfolding comp_def by auto - -lemma mmap: -assumes "f \ multiset" -shows "mmap h f \ multiset" -using assms unfolding mmap_def[abs_def] multiset_def proof safe - assume fin: "finite {a. 0 < f a}" (is "finite ?A") - show "finite {b. 0 < setsum f {a. h a = b \ 0 < f a}}" - (is "finite {b. 0 < setsum f (?As b)}") - proof- let ?B = "{b. 0 < setsum f (?As b)}" - have "\ b. finite (?As b)" using assms unfolding multiset_def by simp - hence B: "?B = {b. ?As b \ {}}" using setsum_gt_0_iff by auto - hence "?B \ h ` ?A" by auto - thus ?thesis using finite_surj[OF fin] by auto - qed -qed - -lemma mmap_cong: -assumes "\a. a \# M \ f a = g a" -shows "mmap f (count M) = mmap g (count M)" -using assms unfolding mmap_def[abs_def] by (intro ext, intro setsum_cong) auto - -abbreviation supp where "supp f \ {a. f a > 0}" - -lemma mmap_image_comp: -assumes f: "f \ multiset" -shows "(supp o mmap h) f = (image h o supp) f" -unfolding mmap_def[abs_def] comp_def proof- - have "\ b. finite {a. h a = b \ 0 < f a}" (is "\ b. finite (?As b)") - using f unfolding multiset_def by auto - thus "{b. 0 < setsum f (?As b)} = h ` {a. 0 < f a}" - using setsum_gt_0_iff by auto -qed - -lemma mmap_image: -assumes f: "f \ multiset" -shows "supp (mmap h f) = h ` (supp f)" -using mmap_image_comp[OF assms] unfolding comp_def . - -lemma set_of_Abs_multiset: -assumes f: "f \ multiset" -shows "set_of (Abs_multiset f) = supp f" -using assms unfolding set_of_def by (auto simp: Abs_multiset_inverse) - -lemma supp_count: -"supp (count M) = set_of M" -using assms unfolding set_of_def by auto - -lemma multiset_of_surj: -"multiset_of ` {as. set as \ A} = {M. set_of M \ A}" -proof safe - fix M assume M: "set_of M \ A" - obtain as where eq: "M = multiset_of as" using surj_multiset_of unfolding surj_def by auto - hence "set as \ A" using M by auto - thus "M \ multiset_of ` {as. set as \ A}" using eq by auto -next - show "\x xa xb. \set xa \ A; xb \ set_of (multiset_of xa)\ \ xb \ A" - by (erule set_mp) (unfold set_of_multiset_of) -qed - -lemma card_of_set_of: -"|{M. set_of M \ A}| \o |{as. set as \ A}|" -apply(rule card_of_ordLeqI2[of _ multiset_of]) using multiset_of_surj by auto - -lemma nat_sum_induct: -assumes "\n1 n2. (\ m1 m2. m1 + m2 < n1 + n2 \ phi m1 m2) \ phi n1 n2" -shows "phi (n1::nat) (n2::nat)" -proof- - let ?chi = "\ n1n2 :: nat * nat. phi (fst n1n2) (snd n1n2)" - have "?chi (n1,n2)" - apply(induct rule: measure_induct[of "\ n1n2. fst n1n2 + snd n1n2" ?chi]) - using assms by (metis fstI sndI) - thus ?thesis by simp -qed - -lemma matrix_count: -fixes ct1 ct2 :: "nat \ nat" -assumes "setsum ct1 {.. ct. (\ i1 \ n1. setsum (\ i2. ct i1 i2) {.. - (\ i2 \ n2. setsum (\ i1. ct i1 i2) {.. ct1 ct2 :: nat \ nat. - setsum ct1 {.. ?phi ct1 ct2 n1 n2" - proof(induct rule: nat_sum_induct[of -"\ n1 n2. \ ct1 ct2 :: nat \ nat. - setsum ct1 {.. ?phi ct1 ct2 n1 n2"], - clarify) - fix n1 n2 :: nat and ct1 ct2 :: "nat \ nat" - assume IH: "\ m1 m2. m1 + m2 < n1 + n2 \ - \ dt1 dt2 :: nat \ nat. - setsum dt1 {.. ?phi dt1 dt2 m1 m2" - and ss: "setsum ct1 {.. ct2 n2") - case True - def dt2 \ "\ i2. if i2 = n2 then ct2 i2 - ct1 n1 else ct2 i2" - have "setsum ct1 {.. i1. i1 \ m1 \ setsum (\ i2. dt i1 i2) {.. i2. i2 \ n2 \ setsum (\ i1. dt i1 i2) {.. "\ i1. if i1 = n1 then ct1 i1 - ct2 n2 else ct1 i1" - have "setsum dt1 {.. i1. i1 \ n1 \ setsum (\ i2. dt i1 i2) {.. i2. i2 \ m2 \ setsum (\ i1. dt i1 i2) {.. - \ b1 b1' b2 b2'. {b1,b1'} \ B1 \ {b2,b2'} \ B2 \ u b1 b2 = u b1' b2' - \ b1 = b1' \ b2 = b2'" - -lemma matrix_count_finite: -assumes B1: "B1 \ {}" "finite B1" and B2: "B2 \ {}" "finite B2" and u: "inj2 u B1 B2" -and ss: "setsum N1 B1 = setsum N2 B2" -shows "\ M :: 'a \ nat. - (\ b1 \ B1. setsum (\ b2. M (u b1 b2)) B2 = N1 b1) \ - (\ b2 \ B2. setsum (\ b1. M (u b1 b2)) B1 = N2 b2)" -proof- - obtain n1 where "card B1 = Suc n1" using B1 by (metis card_insert finite.simps) - then obtain e1 where e1: "bij_betw e1 {.. "inv_into {.. i1. i1 < Suc n1 \ f1 (e1 i1) = i1" - and e1f1[simp]: "\ b1. b1 \ B1 \ e1 (f1 b1) = b1" unfolding f1_def - apply (metis bij_betw_inv_into e1, metis bij_betw_inv_into_left e1 lessThan_iff) - by (metis e1_surj f_inv_into_f) - (* *) - obtain n2 where "card B2 = Suc n2" using B2 by (metis card_insert finite.simps) - then obtain e2 where e2: "bij_betw e2 {.. "inv_into {.. i2. i2 < Suc n2 \ f2 (e2 i2) = i2" - and e2f2[simp]: "\ b2. b2 \ B2 \ e2 (f2 b2) = b2" unfolding f2_def - apply (metis bij_betw_inv_into e2, metis bij_betw_inv_into_left e2 lessThan_iff) - by (metis e2_surj f_inv_into_f) - (* *) - let ?ct1 = "N1 o e1" let ?ct2 = "N2 o e2" - have ss: "setsum ?ct1 {.. i1. i1 \ n1 \ setsum (\ i2. ct i1 i2) {.. i2. i2 \ n2 \ setsum (\ i1. ct i1 i2) {.. "{u b1 b2 | b1 b2. b1 \ B1 \ b2 \ B2}" - have "\ a \ A. \ b1b2 \ B1 <*> B2. u (fst b1b2) (snd b1b2) = a" - unfolding A_def Ball_def mem_Collect_eq by auto - then obtain h1h2 where h12: - "\a. a \ A \ u (fst (h1h2 a)) (snd (h1h2 a)) = a \ h1h2 a \ B1 <*> B2" by metis - def h1 \ "fst o h1h2" def h2 \ "snd o h1h2" - have h12[simp]: "\a. a \ A \ u (h1 a) (h2 a) = a" - "\ a. a \ A \ h1 a \ B1" "\ a. a \ A \ h2 a \ B2" - using h12 unfolding h1_def h2_def by force+ - {fix b1 b2 assume b1: "b1 \ B1" and b2: "b2 \ B2" - hence inA: "u b1 b2 \ A" unfolding A_def by auto - hence "u b1 b2 = u (h1 (u b1 b2)) (h2 (u b1 b2))" by auto - moreover have "h1 (u b1 b2) \ B1" "h2 (u b1 b2) \ B2" using inA by auto - ultimately have "h1 (u b1 b2) = b1 \ h2 (u b1 b2) = b2" - using u b1 b2 unfolding inj2_def by fastforce - } - hence h1[simp]: "\ b1 b2. \b1 \ B1; b2 \ B2\ \ h1 (u b1 b2) = b1" and - h2[simp]: "\ b1 b2. \b1 \ B1; b2 \ B2\ \ h2 (u b1 b2) = b2" by auto - def M \ "\ a. ct (f1 (h1 a)) (f2 (h2 a))" - show ?thesis - apply(rule exI[of _ M]) proof safe - fix b1 assume b1: "b1 \ B1" - hence f1b1: "f1 b1 \ n1" using f1 unfolding bij_betw_def - by (metis bij_betwE f1 lessThan_iff less_Suc_eq_le) - have "(\b2\B2. M (u b1 b2)) = (\i2b2\B2. M (u b1 b2)) = N1 b1" . - next - fix b2 assume b2: "b2 \ B2" - hence f2b2: "f2 b2 \ n2" using f2 unfolding bij_betw_def - by (metis bij_betwE f2 lessThan_iff less_Suc_eq_le) - have "(\b1\B1. M (u b1 b2)) = (\i1b1\B1. M (u b1 b2)) = N2 b2" . - qed -qed - -lemma supp_vimage_mmap: -assumes "M \ multiset" -shows "supp M \ f -` (supp (mmap f M))" -using assms by (auto simp: mmap_image) - -lemma mmap_ge_0: -assumes "M \ multiset" -shows "0 < mmap f M b \ (\a. 0 < M a \ f a = b)" -proof- - have f: "finite {a. f a = b \ 0 < M a}" using assms unfolding multiset_def by auto - show ?thesis unfolding mmap_def setsum_gt_0_iff[OF f] by auto -qed - -lemma finite_twosets: -assumes "finite B1" and "finite B2" -shows "finite {u b1 b2 |b1 b2. b1 \ B1 \ b2 \ B2}" (is "finite ?A") -proof- - have A: "?A = (\ b1b2. u (fst b1b2) (snd b1b2)) ` (B1 <*> B2)" by force - show ?thesis unfolding A using finite_cartesian_product[OF assms] by auto -qed - -lemma wp_mmap: -fixes A :: "'a set" and B1 :: "'b1 set" and B2 :: "'b2 set" -assumes wp: "wpull A B1 B2 f1 f2 p1 p2" -shows -"wpull {M. M \ multiset \ supp M \ A} - {N1. N1 \ multiset \ supp N1 \ B1} {N2. N2 \ multiset \ supp N2 \ B2} - (mmap f1) (mmap f2) (mmap p1) (mmap p2)" -unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq) - fix N1 :: "'b1 \ nat" and N2 :: "'b2 \ nat" - assume mmap': "mmap f1 N1 = mmap f2 N2" - and N1[simp]: "N1 \ multiset" "supp N1 \ B1" - and N2[simp]: "N2 \ multiset" "supp N2 \ B2" - have mN1[simp]: "mmap f1 N1 \ multiset" using N1 by (auto simp: mmap) - have mN2[simp]: "mmap f2 N2 \ multiset" using N2 by (auto simp: mmap) - def P \ "mmap f1 N1" - have P1: "P = mmap f1 N1" and P2: "P = mmap f2 N2" unfolding P_def using mmap' by auto - note P = P1 P2 - have P_mult[simp]: "P \ multiset" unfolding P_def using N1 by auto - have fin_N1[simp]: "finite (supp N1)" using N1(1) unfolding multiset_def by auto - have fin_N2[simp]: "finite (supp N2)" using N2(1) unfolding multiset_def by auto - have fin_P[simp]: "finite (supp P)" using P_mult unfolding multiset_def by auto - (* *) - def set1 \ "\ c. {b1 \ supp N1. f1 b1 = c}" - have set1[simp]: "\ c b1. b1 \ set1 c \ f1 b1 = c" unfolding set1_def by auto - have fin_set1: "\ c. c \ supp P \ finite (set1 c)" - using N1(1) unfolding set1_def multiset_def by auto - have set1_NE: "\ c. c \ supp P \ set1 c \ {}" - unfolding set1_def P1 mmap_ge_0[OF N1(1)] by auto - have supp_N1_set1: "supp N1 = (\ c \ supp P. set1 c)" - using supp_vimage_mmap[OF N1(1), of f1] unfolding set1_def P1 by auto - hence set1_inclN1: "\c. c \ supp P \ set1 c \ supp N1" by auto - hence set1_incl: "\ c. c \ supp P \ set1 c \ B1" using N1(2) by blast - have set1_disj: "\ c c'. c \ c' \ set1 c \ set1 c' = {}" - unfolding set1_def by auto - have setsum_set1: "\ c. setsum N1 (set1 c) = P c" - unfolding P1 set1_def mmap_def apply(rule setsum_cong) by auto - (* *) - def set2 \ "\ c. {b2 \ supp N2. f2 b2 = c}" - have set2[simp]: "\ c b2. b2 \ set2 c \ f2 b2 = c" unfolding set2_def by auto - have fin_set2: "\ c. c \ supp P \ finite (set2 c)" - using N2(1) unfolding set2_def multiset_def by auto - have set2_NE: "\ c. c \ supp P \ set2 c \ {}" - unfolding set2_def P2 mmap_ge_0[OF N2(1)] by auto - have supp_N2_set2: "supp N2 = (\ c \ supp P. set2 c)" - using supp_vimage_mmap[OF N2(1), of f2] unfolding set2_def P2 by auto - hence set2_inclN2: "\c. c \ supp P \ set2 c \ supp N2" by auto - hence set2_incl: "\ c. c \ supp P \ set2 c \ B2" using N2(2) by blast - have set2_disj: "\ c c'. c \ c' \ set2 c \ set2 c' = {}" - unfolding set2_def by auto - have setsum_set2: "\ c. setsum N2 (set2 c) = P c" - unfolding P2 set2_def mmap_def apply(rule setsum_cong) by auto - (* *) - have ss: "\ c. c \ supp P \ setsum N1 (set1 c) = setsum N2 (set2 c)" - unfolding setsum_set1 setsum_set2 .. - have "\ c \ supp P. \ b1b2 \ (set1 c) \ (set2 c). - \ a \ A. p1 a = fst b1b2 \ p2 a = snd b1b2" - using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq - by simp (metis set1 set2 set_rev_mp) - then obtain uu where uu: - "\ c \ supp P. \ b1b2 \ (set1 c) \ (set2 c). - uu c b1b2 \ A \ p1 (uu c b1b2) = fst b1b2 \ p2 (uu c b1b2) = snd b1b2" by metis - def u \ "\ c b1 b2. uu c (b1,b2)" - have u[simp]: - "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ u c b1 b2 \ A" - "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ p1 (u c b1 b2) = b1" - "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ p2 (u c b1 b2) = b2" - using uu unfolding u_def by auto - {fix c assume c: "c \ supp P" - have "inj2 (u c) (set1 c) (set2 c)" unfolding inj2_def proof clarify - fix b1 b1' b2 b2' - assume "{b1, b1'} \ set1 c" "{b2, b2'} \ set2 c" and 0: "u c b1 b2 = u c b1' b2'" - hence "p1 (u c b1 b2) = b1 \ p2 (u c b1 b2) = b2 \ - p1 (u c b1' b2') = b1' \ p2 (u c b1' b2') = b2'" - using u(2)[OF c] u(3)[OF c] by simp metis - thus "b1 = b1' \ b2 = b2'" using 0 by auto - qed - } note inj = this - def sset \ "\ c. {u c b1 b2 | b1 b2. b1 \ set1 c \ b2 \ set2 c}" - have fin_sset[simp]: "\ c. c \ supp P \ finite (sset c)" unfolding sset_def - using fin_set1 fin_set2 finite_twosets by blast - have sset_A: "\ c. c \ supp P \ sset c \ A" unfolding sset_def by auto - {fix c a assume c: "c \ supp P" and ac: "a \ sset c" - then obtain b1 b2 where b1: "b1 \ set1 c" and b2: "b2 \ set2 c" - and a: "a = u c b1 b2" unfolding sset_def by auto - have "p1 a \ set1 c" and p2a: "p2 a \ set2 c" - using ac a b1 b2 c u(2) u(3) by simp+ - hence "u c (p1 a) (p2 a) = a" unfolding a using b1 b2 inj[OF c] - unfolding inj2_def by (metis c u(2) u(3)) - } note u_p12[simp] = this - {fix c a assume c: "c \ supp P" and ac: "a \ sset c" - hence "p1 a \ set1 c" unfolding sset_def by auto - }note p1[simp] = this - {fix c a assume c: "c \ supp P" and ac: "a \ sset c" - hence "p2 a \ set2 c" unfolding sset_def by auto - }note p2[simp] = this - (* *) - {fix c assume c: "c \ supp P" - hence "\ M. (\ b1 \ set1 c. setsum (\ b2. M (u c b1 b2)) (set2 c) = N1 b1) \ - (\ b2 \ set2 c. setsum (\ b1. M (u c b1 b2)) (set1 c) = N2 b2)" - unfolding sset_def - using matrix_count_finite[OF set1_NE[OF c] fin_set1[OF c] - set2_NE[OF c] fin_set2[OF c] inj[OF c] ss[OF c]] by auto - } - then obtain Ms where - ss1: "\ c b1. \c \ supp P; b1 \ set1 c\ \ - setsum (\ b2. Ms c (u c b1 b2)) (set2 c) = N1 b1" and - ss2: "\ c b2. \c \ supp P; b2 \ set2 c\ \ - setsum (\ b1. Ms c (u c b1 b2)) (set1 c) = N2 b2" - by metis - def SET \ "\ c \ supp P. sset c" - have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto - have SET_A: "SET \ A" unfolding SET_def using sset_A by auto - have u_SET[simp]: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ u c b1 b2 \ SET" - unfolding SET_def sset_def by blast - {fix c a assume c: "c \ supp P" and a: "a \ SET" and p1a: "p1 a \ set1 c" - then obtain c' where c': "c' \ supp P" and ac': "a \ sset c'" - unfolding SET_def by auto - hence "p1 a \ set1 c'" unfolding sset_def by auto - hence eq: "c = c'" using p1a c c' set1_disj by auto - hence "a \ sset c" using ac' by simp - } note p1_rev = this - {fix c a assume c: "c \ supp P" and a: "a \ SET" and p2a: "p2 a \ set2 c" - then obtain c' where c': "c' \ supp P" and ac': "a \ sset c'" - unfolding SET_def by auto - hence "p2 a \ set2 c'" unfolding sset_def by auto - hence eq: "c = c'" using p2a c c' set2_disj by auto - hence "a \ sset c" using ac' by simp - } note p2_rev = this - (* *) - have "\ a \ SET. \ c \ supp P. a \ sset c" unfolding SET_def by auto - then obtain h where h: "\ a \ SET. h a \ supp P \ a \ sset (h a)" by metis - have h_u[simp]: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ - \ h (u c b1 b2) = c" - by (metis h p2 set2 u(3) u_SET) - have h_u1: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ - \ h (u c b1 b2) = f1 b1" - using h unfolding sset_def by auto - have h_u2: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ - \ h (u c b1 b2) = f2 b2" - using h unfolding sset_def by auto - def M \ "\ a. if a \ SET \ p1 a \ supp N1 \ p2 a \ supp N2 then Ms (h a) a else 0" - have sM: "supp M \ SET" "supp M \ p1 -` (supp N1)" "supp M \ p2 -` (supp N2)" - unfolding M_def by auto - show "\M. (M \ multiset \ supp M \ A) \ mmap p1 M = N1 \ mmap p2 M = N2" - proof(rule exI[of _ M], safe) - show "M \ multiset" - unfolding multiset_def using finite_subset[OF sM(1) fin_SET] by simp - next - fix a assume "0 < M a" - thus "a \ A" unfolding M_def using SET_A by (cases "a \ SET") auto - next - show "mmap p1 M = N1" - unfolding mmap_def[abs_def] proof(rule ext) - fix b1 - let ?K = "{a. p1 a = b1 \ 0 < M a}" - show "setsum M ?K = N1 b1" - proof(cases "b1 \ supp N1") - case False - hence "?K = {}" using sM(2) by auto - thus ?thesis using False by auto - next - case True - def c \ "f1 b1" - have c: "c \ supp P" and b1: "b1 \ set1 c" - unfolding set1_def c_def P1 using True by (auto simp: mmap_image) - have "setsum M ?K = setsum M {a. p1 a = b1 \ a \ SET}" - apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto - also have "... = setsum M ((\ b2. u c b1 b2) ` (set2 c))" - apply(rule setsum_cong) using c b1 proof safe - fix a assume p1a: "p1 a \ set1 c" and "0 < P c" and "a \ SET" - hence ac: "a \ sset c" using p1_rev by auto - hence "a = u c (p1 a) (p2 a)" using c by auto - moreover have "p2 a \ set2 c" using ac c by auto - ultimately show "a \ u c (p1 a) ` set2 c" by auto - next - fix b2 assume b1: "b1 \ set1 c" and b2: "b2 \ set2 c" - hence "u c b1 b2 \ SET" using c by auto - qed auto - also have "... = setsum (\ b2. M (u c b1 b2)) (set2 c)" - unfolding comp_def[symmetric] apply(rule setsum_reindex) - using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast - also have "... = N1 b1" unfolding ss1[OF c b1, symmetric] - apply(rule setsum_cong[OF refl]) unfolding M_def - using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1] by fastforce - finally show ?thesis . - qed - qed - next - show "mmap p2 M = N2" - unfolding mmap_def[abs_def] proof(rule ext) - fix b2 - let ?K = "{a. p2 a = b2 \ 0 < M a}" - show "setsum M ?K = N2 b2" - proof(cases "b2 \ supp N2") - case False - hence "?K = {}" using sM(3) by auto - thus ?thesis using False by auto - next - case True - def c \ "f2 b2" - have c: "c \ supp P" and b2: "b2 \ set2 c" - unfolding set2_def c_def P2 using True by (auto simp: mmap_image) - have "setsum M ?K = setsum M {a. p2 a = b2 \ a \ SET}" - apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto - also have "... = setsum M ((\ b1. u c b1 b2) ` (set1 c))" - apply(rule setsum_cong) using c b2 proof safe - fix a assume p2a: "p2 a \ set2 c" and "0 < P c" and "a \ SET" - hence ac: "a \ sset c" using p2_rev by auto - hence "a = u c (p1 a) (p2 a)" using c by auto - moreover have "p1 a \ set1 c" using ac c by auto - ultimately show "a \ (\b1. u c b1 (p2 a)) ` set1 c" by auto - next - fix b2 assume b1: "b1 \ set1 c" and b2: "b2 \ set2 c" - hence "u c b1 b2 \ SET" using c by auto - qed auto - also have "... = setsum (M o (\ b1. u c b1 b2)) (set1 c)" - apply(rule setsum_reindex) - using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast - also have "... = setsum (\ b1. M (u c b1 b2)) (set1 c)" - unfolding comp_def[symmetric] by simp - also have "... = N2 b2" unfolding ss2[OF c b2, symmetric] - apply(rule setsum_cong[OF refl]) unfolding M_def set2_def - using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2] - unfolding set1_def by fastforce - finally show ?thesis . - qed - qed - qed -qed - -definition mset_map :: "('a \ 'b) \ 'a multiset \ 'b multiset" where -"mset_map h = Abs_multiset \ mmap h \ count" - -bnf_def mset = mset_map [set_of] "\_::'a multiset. natLeq" ["{#}"] -unfolding mset_map_def -proof - - show "Abs_multiset \ mmap id \ count = id" unfolding mmap_id by (auto simp: count_inverse) -next - fix f g - show "Abs_multiset \ mmap (g \ f) \ count = - Abs_multiset \ mmap g \ count \ (Abs_multiset \ mmap f \ count)" - unfolding comp_def apply(rule ext) - by (auto simp: Abs_multiset_inverse count mmap_comp1 mmap) -next - fix M f g assume eq: "\a. a \ set_of M \ f a = g a" - thus "(Abs_multiset \ mmap f \ count) M = (Abs_multiset \ mmap g \ count) M" apply auto - unfolding cIm_def[abs_def] image_def - by (auto intro!: mmap_cong simp: Abs_multiset_inject count mmap) -next - fix f show "set_of \ (Abs_multiset \ mmap f \ count) = op ` f \ set_of" - by (auto simp: count mmap mmap_image set_of_Abs_multiset supp_count) -next - show "card_order natLeq" by (rule natLeq_card_order) -next - show "cinfinite natLeq" by (rule natLeq_cinfinite) -next - fix M show "|set_of M| \o natLeq" - apply(rule ordLess_imp_ordLeq) - unfolding finite_iff_ordLess_natLeq[symmetric] using finite_set_of . -next - fix A :: "'a set" - have "|{M. set_of M \ A}| \o |{as. set as \ A}|" using card_of_set_of . - also have "|{as. set as \ A}| \o ( |A| +c ctwo) ^c natLeq" - by (rule list_in_bd) - finally show "|{M. set_of M \ A}| \o ( |A| +c ctwo) ^c natLeq" . -next - fix A B1 B2 f1 f2 p1 p2 - let ?map = "\ f. Abs_multiset \ mmap f \ count" - assume wp: "wpull A B1 B2 f1 f2 p1 p2" - show "wpull {x. set_of x \ A} {x. set_of x \ B1} {x. set_of x \ B2} - (?map f1) (?map f2) (?map p1) (?map p2)" - unfolding wpull_def proof safe - fix y1 y2 - assume y1: "set_of y1 \ B1" and y2: "set_of y2 \ B2" - and m: "?map f1 y1 = ?map f2 y2" - def N1 \ "count y1" def N2 \ "count y2" - have "N1 \ multiset \ supp N1 \ B1" and "N2 \ multiset \ supp N2 \ B2" - and "mmap f1 N1 = mmap f2 N2" - using y1 y2 m unfolding N1_def N2_def - by (auto simp: Abs_multiset_inject count mmap) - then obtain M where M: "M \ multiset \ supp M \ A" - and N1: "mmap p1 M = N1" and N2: "mmap p2 M = N2" - using wp_mmap[OF wp] unfolding wpull_def by auto - def x \ "Abs_multiset M" - show "\x\{x. set_of x \ A}. ?map p1 x = y1 \ ?map p2 x = y2" - apply(intro bexI[of _ x]) using M N1 N2 unfolding N1_def N2_def x_def - by (auto simp: count_inverse Abs_multiset_inverse) - qed -qed (unfold set_of_empty, auto) - end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/Codatatype.thy --- a/src/HOL/Codatatype/Codatatype.thy Wed Sep 12 05:03:18 2012 +0200 +++ b/src/HOL/Codatatype/Codatatype.thy Wed Sep 12 05:21:47 2012 +0200 @@ -10,7 +10,7 @@ header {* The (Co)datatype Package *} theory Codatatype -imports BNF_LFP BNF_GFP +imports More_BNFs begin end diff -r 6190b701e4f4 -r f20b24214ac2 src/HOL/Codatatype/More_BNFs.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Codatatype/More_BNFs.thy Wed Sep 12 05:21:47 2012 +0200 @@ -0,0 +1,1180 @@ +(* Title: HOL/Codatatype/More_BNFs.thy + Author: Dmitriy Traytel, TU Muenchen + Author: Andrei Popescu, TU Muenchen + Author: Andreas Lochbihler, Karlsruhe Institute of Technology + Author: Jasmin Blanchette, TU Muenchen + Copyright 2012 + +Registration of various types as bounded natural functors. +*) + +header {* Registration of Various Types as Bounded Natural Functors *} + +theory More_BNFs +imports BNF_LFP BNF_GFP +begin + +lemma option_rec_conv_option_case: "option_rec = option_case" +by (simp add: fun_eq_iff split: option.split) + +bnf_def option = Option.map [Option.set] "\_::'a option. natLeq" ["None"] +proof - + show "Option.map id = id" by (simp add: fun_eq_iff Option.map_def split: option.split) +next + fix f g + show "Option.map (g \ f) = Option.map g \ Option.map f" + by (auto simp add: fun_eq_iff Option.map_def split: option.split) +next + fix f g x + assume "\z. z \ Option.set x \ f z = g z" + thus "Option.map f x = Option.map g x" + by (simp cong: Option.map_cong) +next + fix f + show "Option.set \ Option.map f = op ` f \ Option.set" + by fastforce +next + show "card_order natLeq" by (rule natLeq_card_order) +next + show "cinfinite natLeq" by (rule natLeq_cinfinite) +next + fix x + show "|Option.set x| \o natLeq" + by (cases x) (simp_all add: ordLess_imp_ordLeq finite_iff_ordLess_natLeq[symmetric]) +next + fix A + have unfold: "{x. Option.set x \ A} = Some ` A \ {None}" + by (auto simp add: option_rec_conv_option_case Option.set_def split: option.split_asm) + show "|{x. Option.set x \ A}| \o ( |A| +c ctwo) ^c natLeq" + apply (rule ordIso_ordLeq_trans) + apply (rule card_of_ordIso_subst[OF unfold]) + apply (rule ordLeq_transitive) + apply (rule Un_csum) + apply (rule ordLeq_transitive) + apply (rule csum_mono) + apply (rule card_of_image) + apply (rule ordIso_ordLeq_trans) + apply (rule single_cone) + apply (rule cone_ordLeq_ctwo) + apply (rule ordLeq_cexp1) + apply (simp_all add: natLeq_cinfinite natLeq_Card_order cinfinite_not_czero Card_order_csum) + done +next + fix A B1 B2 f1 f2 p1 p2 + assume wpull: "wpull A B1 B2 f1 f2 p1 p2" + show "wpull {x. Option.set x \ A} {x. Option.set x \ B1} {x. Option.set x \ B2} + (Option.map f1) (Option.map f2) (Option.map p1) (Option.map p2)" + (is "wpull ?A ?B1 ?B2 ?f1 ?f2 ?p1 ?p2") + unfolding wpull_def + proof (intro strip, elim conjE) + fix b1 b2 + assume "b1 \ ?B1" "b2 \ ?B2" "?f1 b1 = ?f2 b2" + thus "\a \ ?A. ?p1 a = b1 \ ?p2 a = b2" using wpull + unfolding wpull_def by (cases b2) (auto 4 5) + qed +next + fix z + assume "z \ Option.set None" + thus False by simp +qed + +lemma card_of_list_in: + "|{xs. set xs \ A}| \o |Pfunc (UNIV :: nat set) A|" (is "|?LHS| \o |?RHS|") +proof - + let ?f = "%xs. %i. if i < length xs \ set xs \ A then Some (nth xs i) else None" + have "inj_on ?f ?LHS" unfolding inj_on_def fun_eq_iff + proof safe + fix xs :: "'a list" and ys :: "'a list" + assume su: "set xs \ A" "set ys \ A" and eq: "\i. ?f xs i = ?f ys i" + hence *: "length xs = length ys" + by (metis linorder_cases option.simps(2) order_less_irrefl) + thus "xs = ys" by (rule nth_equalityI) (metis * eq su option.inject) + qed + moreover have "?f ` ?LHS \ ?RHS" unfolding Pfunc_def by fastforce + ultimately show ?thesis using card_of_ordLeq by blast +qed + +lemma list_in_empty: "A = {} \ {x. set x \ A} = {[]}" +by simp + +lemma card_of_Func: "|Func A B| =o |B| ^c |A|" +unfolding cexp_def Field_card_of by (rule card_of_refl) + +lemma not_emp_czero_notIn_ordIso_Card_order: +"A \ {} \ ( |A|, czero) \ ordIso \ Card_order |A|" + apply (rule conjI) + apply (metis Field_card_of czeroE) + by (rule card_of_Card_order) + +lemma list_in_bd: "|{x. set x \ A}| \o ( |A| +c ctwo) ^c natLeq" +proof - + fix A :: "'a set" + show "|{x. set x \ A}| \o ( |A| +c ctwo) ^c natLeq" + proof (cases "A = {}") + case False thus ?thesis + apply - + apply (rule ordLeq_transitive) + apply (rule card_of_list_in) + apply (rule ordLeq_transitive) + apply (erule card_of_Pfunc_Pow_Func) + apply (rule ordIso_ordLeq_trans) + apply (rule Times_cprod) + apply (rule cprod_cinfinite_bound) + apply (rule ordIso_ordLeq_trans) + apply (rule Pow_cexp_ctwo) + apply (rule ordIso_ordLeq_trans) + apply (rule cexp_cong2) + apply (rule card_of_nat) + apply (rule Card_order_ctwo) + apply (rule card_of_Card_order) + apply (rule natLeq_Card_order) + apply (rule disjI1) + apply (rule ctwo_Cnotzero) + apply (rule cexp_mono1) + apply (rule ordLeq_csum2) + apply (rule Card_order_ctwo) + apply (rule disjI1) + apply (rule ctwo_Cnotzero) + apply (rule natLeq_Card_order) + apply (rule ordIso_ordLeq_trans) + apply (rule card_of_Func) + apply (rule ordIso_ordLeq_trans) + apply (rule cexp_cong2) + apply (rule card_of_nat) + apply (rule card_of_Card_order) + apply (rule card_of_Card_order) + apply (rule natLeq_Card_order) + apply (rule disjI1) + apply (erule not_emp_czero_notIn_ordIso_Card_order) + apply (rule cexp_mono1) + apply (rule ordLeq_csum1) + apply (rule card_of_Card_order) + apply (rule disjI1) + apply (erule not_emp_czero_notIn_ordIso_Card_order) + apply (rule natLeq_Card_order) + apply (rule card_of_Card_order) + apply (rule card_of_Card_order) + apply (rule Cinfinite_cexp) + apply (rule ordLeq_csum2) + apply (rule Card_order_ctwo) + apply (rule conjI) + apply (rule natLeq_cinfinite) + by (rule natLeq_Card_order) + next + case True thus ?thesis + apply - + apply (rule ordIso_ordLeq_trans) + apply (rule card_of_ordIso_subst) + apply (erule list_in_empty) + apply (rule ordIso_ordLeq_trans) + apply (rule single_cone) + apply (rule cone_ordLeq_cexp) + apply (rule ordLeq_transitive) + apply (rule cone_ordLeq_ctwo) + apply (rule ordLeq_csum2) + by (rule Card_order_ctwo) + qed +qed + +bnf_def list = map [set] "\_::'a list. natLeq" ["[]"] +proof - + show "map id = id" by (rule List.map.id) +next + fix f g + show "map (g o f) = map g o map f" by (rule List.map.comp[symmetric]) +next + fix x f g + assume "\z. z \ set x \ f z = g z" + thus "map f x = map g x" by simp +next + fix f + show "set o map f = image f o set" by (rule ext, unfold o_apply, rule set_map) +next + show "card_order natLeq" by (rule natLeq_card_order) +next + show "cinfinite natLeq" by (rule natLeq_cinfinite) +next + fix x + show "|set x| \o natLeq" + apply (rule ordLess_imp_ordLeq) + apply (rule finite_ordLess_infinite[OF _ natLeq_Well_order]) + unfolding Field_natLeq Field_card_of by (auto simp: card_of_well_order_on) +next + fix A :: "'a set" + show "|{x. set x \ A}| \o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd) +next + fix A B1 B2 f1 f2 p1 p2 + assume "wpull A B1 B2 f1 f2 p1 p2" + hence pull: "\b1 b2. b1 \ B1 \ b2 \ B2 \ f1 b1 = f2 b2 \ \a \ A. p1 a = b1 \ p2 a = b2" + unfolding wpull_def by auto + show "wpull {x. set x \ A} {x. set x \ B1} {x. set x \ B2} (map f1) (map f2) (map p1) (map p2)" + (is "wpull ?A ?B1 ?B2 _ _ _ _") + proof (unfold wpull_def) + { fix as bs assume *: "as \ ?B1" "bs \ ?B2" "map f1 as = map f2 bs" + hence "length as = length bs" by (metis length_map) + hence "\zs \ ?A. map p1 zs = as \ map p2 zs = bs" using * + proof (induct as bs rule: list_induct2) + case (Cons a as b bs) + hence "a \ B1" "b \ B2" "f1 a = f2 b" by auto + with pull obtain z where "z \ A" "p1 z = a" "p2 z = b" by blast + moreover + from Cons obtain zs where "zs \ ?A" "map p1 zs = as" "map p2 zs = bs" by auto + ultimately have "z # zs \ ?A" "map p1 (z # zs) = a # as \ map p2 (z # zs) = b # bs" by auto + thus ?case by (rule_tac x = "z # zs" in bexI) + qed simp + } + thus "\as bs. as \ ?B1 \ bs \ ?B2 \ map f1 as = map f2 bs \ + (\zs \ ?A. map p1 zs = as \ map p2 zs = bs)" by blast + qed +qed auto + +(* Finite sets *) +abbreviation afset where "afset \ abs_fset" +abbreviation rfset where "rfset \ rep_fset" + +lemma fset_fset_member: +"fset A = {a. a |\| A}" +unfolding fset_def fset_member_def by auto + +lemma afset_rfset: +"afset (rfset x) = x" +by (rule Quotient_fset[unfolded Quotient_def, THEN conjunct1, rule_format]) + +lemma afset_rfset_id: +"afset o rfset = id" +unfolding comp_def afset_rfset id_def .. + +lemma rfset: +"rfset A = rfset B \ A = B" +by (metis afset_rfset) + +lemma afset_set: +"afset as = afset bs \ set as = set bs" +using Quotient_fset unfolding Quotient_def list_eq_def by auto + +lemma surj_afset: +"\ as. A = afset as" +by (metis afset_rfset) + +lemma fset_def2: +"fset = set o rfset" +unfolding fset_def map_fun_def[abs_def] by simp + +lemma fset_def2_raw: +"fset A = set (rfset A)" +unfolding fset_def2 by simp + +lemma fset_comp_afset: +"fset o afset = set" +unfolding fset_def2 comp_def apply(rule ext) +unfolding afset_set[symmetric] afset_rfset .. + +lemma fset_afset: +"fset (afset as) = set as" +unfolding fset_comp_afset[symmetric] by simp + +lemma set_rfset_afset: +"set (rfset (afset as)) = set as" +unfolding afset_set[symmetric] afset_rfset .. + +lemma map_fset_comp_afset: +"(map_fset f) o afset = afset o (map f)" +unfolding map_fset_def map_fun_def[abs_def] comp_def apply(rule ext) +unfolding afset_set set_map set_rfset_afset id_apply .. + +lemma map_fset_afset: +"(map_fset f) (afset as) = afset (map f as)" +using map_fset_comp_afset unfolding comp_def fun_eq_iff by auto + +lemma fset_map_fset: +"fset (map_fset f A) = (image f) (fset A)" +apply(subst afset_rfset[symmetric, of A]) +unfolding map_fset_afset fset_afset set_map +unfolding fset_def2_raw .. + +lemma map_fset_def2: +"map_fset f = afset o (map f) o rfset" +unfolding map_fset_def map_fun_def[abs_def] by simp + +lemma map_fset_def2_raw: +"map_fset f A = afset (map f (rfset A))" +unfolding map_fset_def2 by simp + +lemma finite_ex_fset: +assumes "finite A" +shows "\ B. fset B = A" +by (metis assms finite_list fset_afset) + +lemma wpull_image: +assumes "wpull A B1 B2 f1 f2 p1 p2" +shows "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)" +unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify + fix Y1 Y2 assume Y1: "Y1 \ B1" and Y2: "Y2 \ B2" and EQ: "f1 ` Y1 = f2 ` Y2" + def X \ "{a \ A. p1 a \ Y1 \ p2 a \ Y2}" + show "\X\A. p1 ` X = Y1 \ p2 ` X = Y2" + proof (rule exI[of _ X], intro conjI) + show "p1 ` X = Y1" + proof + show "Y1 \ p1 ` X" + proof safe + fix y1 assume y1: "y1 \ Y1" + then obtain y2 where y2: "y2 \ Y2" and eq: "f1 y1 = f2 y2" using EQ by auto + then obtain x where "x \ A" and "p1 x = y1" and "p2 x = y2" + using assms y1 Y1 Y2 unfolding wpull_def by blast + thus "y1 \ p1 ` X" unfolding X_def using y1 y2 by auto + qed + qed(unfold X_def, auto) + show "p2 ` X = Y2" + proof + show "Y2 \ p2 ` X" + proof safe + fix y2 assume y2: "y2 \ Y2" + then obtain y1 where y1: "y1 \ Y1" and eq: "f1 y1 = f2 y2" using EQ by force + then obtain x where "x \ A" and "p1 x = y1" and "p2 x = y2" + using assms y2 Y1 Y2 unfolding wpull_def by blast + thus "y2 \ p2 ` X" unfolding X_def using y1 y2 by auto + qed + qed(unfold X_def, auto) + qed(unfold X_def, auto) +qed + +lemma fset_to_fset: "finite A \ fset (the_inv fset A) = A" +by (rule f_the_inv_into_f) (auto simp: inj_on_def fset_cong dest!: finite_ex_fset) + +bnf_def fset = map_fset [fset] "\_::'a fset. natLeq" ["{||}"] +proof - + show "map_fset id = id" + unfolding map_fset_def2 map_id o_id afset_rfset_id .. +next + fix f g + show "map_fset (g o f) = map_fset g o map_fset f" + unfolding map_fset_def2 map.comp[symmetric] comp_def apply(rule ext) + unfolding afset_set set_map fset_def2_raw[symmetric] image_image[symmetric] + unfolding map_fset_afset[symmetric] map_fset_image afset_rfset + by (rule refl) +next + fix x f g + assume "\z. z \ fset x \ f z = g z" + hence "map f (rfset x) = map g (rfset x)" + apply(intro map_cong) unfolding fset_def2_raw by auto + thus "map_fset f x = map_fset g x" unfolding map_fset_def2_raw + by (rule arg_cong) +next + fix f + show "fset o map_fset f = image f o fset" + unfolding comp_def fset_map_fset .. +next + show "card_order natLeq" by (rule natLeq_card_order) +next + show "cinfinite natLeq" by (rule natLeq_cinfinite) +next + fix x + show "|fset x| \o natLeq" + unfolding fset_def2_raw + apply (rule ordLess_imp_ordLeq) + apply (rule finite_iff_ordLess_natLeq[THEN iffD1]) + by (rule finite_set) +next + fix A :: "'a set" + have "|{x. fset x \ A}| \o |afset ` {as. set as \ A}|" + apply(rule card_of_mono1) unfolding fset_def2_raw apply auto + apply (rule image_eqI) + by (auto simp: afset_rfset) + also have "|afset ` {as. set as \ A}| \o |{as. set as \ A}|" using card_of_image . + also have "|{as. set as \ A}| \o ( |A| +c ctwo) ^c natLeq" by (rule list_in_bd) + finally show "|{x. fset x \ A}| \o ( |A| +c ctwo) ^c natLeq" . +next + fix A B1 B2 f1 f2 p1 p2 + assume wp: "wpull A B1 B2 f1 f2 p1 p2" + hence "wpull (Pow A) (Pow B1) (Pow B2) (image f1) (image f2) (image p1) (image p2)" + by (rule wpull_image) + show "wpull {x. fset x \ A} {x. fset x \ B1} {x. fset x \ B2} + (map_fset f1) (map_fset f2) (map_fset p1) (map_fset p2)" + unfolding wpull_def Pow_def Bex_def mem_Collect_eq proof clarify + fix y1 y2 + assume Y1: "fset y1 \ B1" and Y2: "fset y2 \ B2" + assume "map_fset f1 y1 = map_fset f2 y2" + hence EQ: "f1 ` (fset y1) = f2 ` (fset y2)" unfolding map_fset_def2_raw + unfolding afset_set set_map fset_def2_raw . + with Y1 Y2 obtain X where X: "X \ A" + and Y1: "p1 ` X = fset y1" and Y2: "p2 ` X = fset y2" + using wpull_image[OF wp] unfolding wpull_def Pow_def + unfolding Bex_def mem_Collect_eq apply - + apply(erule allE[of _ "fset y1"], erule allE[of _ "fset y2"]) by auto + have "\ y1' \ fset y1. \ x. x \ X \ y1' = p1 x" using Y1 by auto + then obtain q1 where q1: "\ y1' \ fset y1. q1 y1' \ X \ y1' = p1 (q1 y1')" by metis + have "\ y2' \ fset y2. \ x. x \ X \ y2' = p2 x" using Y2 by auto + then obtain q2 where q2: "\ y2' \ fset y2. q2 y2' \ X \ y2' = p2 (q2 y2')" by metis + def X' \ "q1 ` (fset y1) \ q2 ` (fset y2)" + have X': "X' \ A" and Y1: "p1 ` X' = fset y1" and Y2: "p2 ` X' = fset y2" + using X Y1 Y2 q1 q2 unfolding X'_def by auto + have fX': "finite X'" unfolding X'_def by simp + then obtain x where X'eq: "X' = fset x" by (auto dest: finite_ex_fset) + show "\x. fset x \ A \ map_fset p1 x = y1 \ map_fset p2 x = y2" + apply(intro exI[of _ "x"]) using X' Y1 Y2 + unfolding X'eq map_fset_def2_raw fset_def2_raw set_map[symmetric] + afset_set[symmetric] afset_rfset by simp + qed +qed auto + +lemma fset_pred[simp]: "fset_pred R a b \ + ((\t \ fset a. (\u \ fset b. R t u)) \ + (\t \ fset b. (\u \ fset a. R u t)))" (is "?L = ?R") +proof + assume ?L thus ?R unfolding fset_rel_def fset_pred_def + Gr_def relcomp_unfold converse_unfold + apply (simp add: subset_eq Ball_def) + apply (rule conjI) + apply (clarsimp, metis snd_conv) + by (clarsimp, metis fst_conv) +next + assume ?R + def R' \ "the_inv fset (Collect (split R) \ (fset a \ fset b))" (is "the_inv fset ?R'") + have "finite ?R'" by (intro finite_Int[OF disjI2] finite_cartesian_product) auto + hence *: "fset R' = ?R'" unfolding R'_def by (intro fset_to_fset) + show ?L unfolding fset_rel_def fset_pred_def Gr_def relcomp_unfold converse_unfold + proof (intro CollectI prod_caseI exI conjI) + from * show "(R', a) = (R', map_fset fst R')" using conjunct1[OF `?R`] + by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits) + from * show "(R', b) = (R', map_fset snd R')" using conjunct2[OF `?R`] + by (auto simp add: fset_cong[symmetric] image_def Int_def split: prod.splits) + qed (auto simp add: *) +qed + +(* Countable sets *) + +lemma card_of_countable_sets_range: +fixes A :: "'a set" +shows "|{X. X \ A \ countable X \ X \ {}}| \o |{f::nat \ 'a. range f \ A}|" +apply(rule card_of_ordLeqI[of fromNat]) using inj_on_fromNat +unfolding inj_on_def by auto + +lemma card_of_countable_sets_Func: +"|{X. X \ A \ countable X \ X \ {}}| \o |A| ^c natLeq" +using card_of_countable_sets_range card_of_Func_UNIV[THEN ordIso_symmetric] +unfolding cexp_def Field_natLeq Field_card_of +by (rule ordLeq_ordIso_trans) + +lemma ordLeq_countable_subsets: +"|A| \o |{X. X \ A \ countable X}|" +apply (rule card_of_ordLeqI[of "\ a. {a}"]) unfolding inj_on_def by auto + +lemma finite_countable_subset: +"finite {X. X \ A \ countable X} \ finite A" +apply default + apply (erule contrapos_pp) + apply (rule card_of_ordLeq_infinite) + apply (rule ordLeq_countable_subsets) + apply assumption +apply (rule finite_Collect_conjI) +apply (rule disjI1) +by (erule finite_Collect_subsets) + +lemma card_of_countable_sets: +"|{X. X \ A \ countable X}| \o ( |A| +c ctwo) ^c natLeq" +(is "|?L| \o _") +proof(cases "finite A") + let ?R = "Func (UNIV::nat set) (A <+> (UNIV::bool set))" + case True hence "finite ?L" by simp + moreover have "infinite ?R" + apply(rule infinite_Func[of _ "Inr True" "Inr False"]) by auto + ultimately show ?thesis unfolding cexp_def csum_def ctwo_def Field_natLeq Field_card_of + apply(intro ordLess_imp_ordLeq) by (rule finite_ordLess_infinite2) +next + case False + hence "|{X. X \ A \ countable X}| =o |{X. X \ A \ countable X} - {{}}|" + by (intro card_of_infinite_diff_finitte finite.emptyI finite.insertI ordIso_symmetric) + (unfold finite_countable_subset) + also have "|{X. X \ A \ countable X} - {{}}| \o |A| ^c natLeq" + using card_of_countable_sets_Func[of A] unfolding set_diff_eq by auto + also have "|A| ^c natLeq \o ( |A| +c ctwo) ^c natLeq" + apply(rule cexp_mono1_cone_ordLeq) + apply(rule ordLeq_csum1, rule card_of_Card_order) + apply (rule cone_ordLeq_cexp) + apply (rule cone_ordLeq_Cnotzero) + using csum_Cnotzero2 ctwo_Cnotzero apply blast + by (rule natLeq_Card_order) + finally show ?thesis . +qed + +bnf_def cset = cIm [rcset] "\_::'a cset. natLeq" ["cEmp"] +proof - + show "cIm id = id" unfolding cIm_def[abs_def] id_def by auto +next + fix f g show "cIm (g \ f) = cIm g \ cIm f" + unfolding cIm_def[abs_def] apply(rule ext) unfolding comp_def by auto +next + fix C f g assume eq: "\a. a \ rcset C \ f a = g a" + thus "cIm f C = cIm g C" + unfolding cIm_def[abs_def] unfolding image_def by auto +next + fix f show "rcset \ cIm f = op ` f \ rcset" unfolding cIm_def[abs_def] by auto +next + show "card_order natLeq" by (rule natLeq_card_order) +next + show "cinfinite natLeq" by (rule natLeq_cinfinite) +next + fix C show "|rcset C| \o natLeq" using rcset unfolding countable_def . +next + fix A :: "'a set" + have "|{Z. rcset Z \ A}| \o |acset ` {X. X \ A \ countable X}|" + apply(rule card_of_mono1) unfolding Pow_def image_def + proof (rule Collect_mono, clarsimp) + fix x + assume "rcset x \ A" + hence "rcset x \ A \ countable (rcset x) \ x = acset (rcset x)" + using acset_rcset[of x] rcset[of x] by force + thus "\y \ A. countable y \ x = acset y" by blast + qed + also have "|acset ` {X. X \ A \ countable X}| \o |{X. X \ A \ countable X}|" + using card_of_image . + also have "|{X. X \ A \ countable X}| \o ( |A| +c ctwo) ^c natLeq" + using card_of_countable_sets . + finally show "|{Z. rcset Z \ A}| \o ( |A| +c ctwo) ^c natLeq" . +next + fix A B1 B2 f1 f2 p1 p2 + assume wp: "wpull A B1 B2 f1 f2 p1 p2" + show "wpull {x. rcset x \ A} {x. rcset x \ B1} {x. rcset x \ B2} + (cIm f1) (cIm f2) (cIm p1) (cIm p2)" + unfolding wpull_def proof safe + fix y1 y2 + assume Y1: "rcset y1 \ B1" and Y2: "rcset y2 \ B2" + assume "cIm f1 y1 = cIm f2 y2" + hence EQ: "f1 ` (rcset y1) = f2 ` (rcset y2)" + unfolding cIm_def by auto + with Y1 Y2 obtain X where X: "X \ A" + and Y1: "p1 ` X = rcset y1" and Y2: "p2 ` X = rcset y2" + using wpull_image[OF wp] unfolding wpull_def Pow_def + unfolding Bex_def mem_Collect_eq apply - + apply(erule allE[of _ "rcset y1"], erule allE[of _ "rcset y2"]) by auto + have "\ y1' \ rcset y1. \ x. x \ X \ y1' = p1 x" using Y1 by auto + then obtain q1 where q1: "\ y1' \ rcset y1. q1 y1' \ X \ y1' = p1 (q1 y1')" by metis + have "\ y2' \ rcset y2. \ x. x \ X \ y2' = p2 x" using Y2 by auto + then obtain q2 where q2: "\ y2' \ rcset y2. q2 y2' \ X \ y2' = p2 (q2 y2')" by metis + def X' \ "q1 ` (rcset y1) \ q2 ` (rcset y2)" + have X': "X' \ A" and Y1: "p1 ` X' = rcset y1" and Y2: "p2 ` X' = rcset y2" + using X Y1 Y2 q1 q2 unfolding X'_def by fast+ + have fX': "countable X'" unfolding X'_def by simp + then obtain x where X'eq: "X' = rcset x" by (metis rcset_acset) + show "\x\{x. rcset x \ A}. cIm p1 x = y1 \ cIm p2 x = y2" + apply(intro bexI[of _ "x"]) using X' Y1 Y2 unfolding X'eq cIm_def by auto + qed +qed (unfold cEmp_def, auto) + + +(* Multisets *) + +lemma setsum_gt_0_iff: +fixes f :: "'a \ nat" assumes "finite A" +shows "setsum f A > 0 \ (\ a \ A. f a > 0)" +(is "?L \ ?R") +proof- + have "?L \ \ setsum f A = 0" by fast + also have "... \ (\ a \ A. f a \ 0)" using assms by simp + also have "... \ ?R" by simp + finally show ?thesis . +qed + +(* *) +definition mmap :: "('a \ 'b) \ ('a \ nat) \ 'b \ nat" where +"mmap h f b = setsum f {a. h a = b \ f a > 0}" + +lemma mmap_id: "mmap id = id" +proof (rule ext)+ + fix f a show "mmap id f a = id f a" + proof(cases "f a = 0") + case False + hence 1: "{aa. aa = a \ 0 < f aa} = {a}" by auto + show ?thesis by (simp add: mmap_def id_apply 1) + qed(unfold mmap_def, auto) +qed + +lemma inj_on_setsum_inv: +assumes f: "f \ multiset" +and 1: "(0::nat) < setsum f {a. h a = b' \ 0 < f a}" (is "0 < setsum f ?A'") +and 2: "{a. h a = b \ 0 < f a} = {a. h a = b' \ 0 < f a}" (is "?A = ?A'") +shows "b = b'" +proof- + have "finite ?A'" using f unfolding multiset_def by auto + hence "?A' \ {}" using 1 setsum_gt_0_iff by auto + thus ?thesis using 2 by auto +qed + +lemma mmap_comp: +fixes h1 :: "'a \ 'b" and h2 :: "'b \ 'c" +assumes f: "f \ multiset" +shows "mmap (h2 o h1) f = (mmap h2 o mmap h1) f" +unfolding mmap_def[abs_def] comp_def proof(rule ext)+ + fix c :: 'c + let ?A = "{a. h2 (h1 a) = c \ 0 < f a}" + let ?As = "\ b. {a. h1 a = b \ 0 < f a}" + let ?B = "{b. h2 b = c \ 0 < setsum f (?As b)}" + have 0: "{?As b | b. b \ ?B} = ?As ` ?B" by auto + have "\ b. finite (?As b)" using f unfolding multiset_def by simp + hence "?B = {b. h2 b = c \ ?As b \ {}}" using setsum_gt_0_iff by auto + hence A: "?A = \ {?As b | b. b \ ?B}" by auto + have "setsum f ?A = setsum (setsum f) {?As b | b. b \ ?B}" + unfolding A apply(rule setsum_Union_disjoint) + using f unfolding multiset_def by auto + also have "... = setsum (setsum f) (?As ` ?B)" unfolding 0 .. + also have "... = setsum (setsum f o ?As) ?B" apply(rule setsum_reindex) + unfolding inj_on_def apply auto using inj_on_setsum_inv[OF f, of h1] by blast + also have "... = setsum (\ b. setsum f (?As b)) ?B" unfolding comp_def .. + finally show "setsum f ?A = setsum (\ b. setsum f (?As b)) ?B" . +qed + +lemma mmap_comp1: +fixes h1 :: "'a \ 'b" and h2 :: "'b \ 'c" +assumes "f \ multiset" +shows "mmap (\ a. h2 (h1 a)) f = mmap h2 (mmap h1 f)" +using mmap_comp[OF assms] unfolding comp_def by auto + +lemma mmap: +assumes "f \ multiset" +shows "mmap h f \ multiset" +using assms unfolding mmap_def[abs_def] multiset_def proof safe + assume fin: "finite {a. 0 < f a}" (is "finite ?A") + show "finite {b. 0 < setsum f {a. h a = b \ 0 < f a}}" + (is "finite {b. 0 < setsum f (?As b)}") + proof- let ?B = "{b. 0 < setsum f (?As b)}" + have "\ b. finite (?As b)" using assms unfolding multiset_def by simp + hence B: "?B = {b. ?As b \ {}}" using setsum_gt_0_iff by auto + hence "?B \ h ` ?A" by auto + thus ?thesis using finite_surj[OF fin] by auto + qed +qed + +lemma mmap_cong: +assumes "\a. a \# M \ f a = g a" +shows "mmap f (count M) = mmap g (count M)" +using assms unfolding mmap_def[abs_def] by (intro ext, intro setsum_cong) auto + +abbreviation supp where "supp f \ {a. f a > 0}" + +lemma mmap_image_comp: +assumes f: "f \ multiset" +shows "(supp o mmap h) f = (image h o supp) f" +unfolding mmap_def[abs_def] comp_def proof- + have "\ b. finite {a. h a = b \ 0 < f a}" (is "\ b. finite (?As b)") + using f unfolding multiset_def by auto + thus "{b. 0 < setsum f (?As b)} = h ` {a. 0 < f a}" + using setsum_gt_0_iff by auto +qed + +lemma mmap_image: +assumes f: "f \ multiset" +shows "supp (mmap h f) = h ` (supp f)" +using mmap_image_comp[OF assms] unfolding comp_def . + +lemma set_of_Abs_multiset: +assumes f: "f \ multiset" +shows "set_of (Abs_multiset f) = supp f" +using assms unfolding set_of_def by (auto simp: Abs_multiset_inverse) + +lemma supp_count: +"supp (count M) = set_of M" +using assms unfolding set_of_def by auto + +lemma multiset_of_surj: +"multiset_of ` {as. set as \ A} = {M. set_of M \ A}" +proof safe + fix M assume M: "set_of M \ A" + obtain as where eq: "M = multiset_of as" using surj_multiset_of unfolding surj_def by auto + hence "set as \ A" using M by auto + thus "M \ multiset_of ` {as. set as \ A}" using eq by auto +next + show "\x xa xb. \set xa \ A; xb \ set_of (multiset_of xa)\ \ xb \ A" + by (erule set_mp) (unfold set_of_multiset_of) +qed + +lemma card_of_set_of: +"|{M. set_of M \ A}| \o |{as. set as \ A}|" +apply(rule card_of_ordLeqI2[of _ multiset_of]) using multiset_of_surj by auto + +lemma nat_sum_induct: +assumes "\n1 n2. (\ m1 m2. m1 + m2 < n1 + n2 \ phi m1 m2) \ phi n1 n2" +shows "phi (n1::nat) (n2::nat)" +proof- + let ?chi = "\ n1n2 :: nat * nat. phi (fst n1n2) (snd n1n2)" + have "?chi (n1,n2)" + apply(induct rule: measure_induct[of "\ n1n2. fst n1n2 + snd n1n2" ?chi]) + using assms by (metis fstI sndI) + thus ?thesis by simp +qed + +lemma matrix_count: +fixes ct1 ct2 :: "nat \ nat" +assumes "setsum ct1 {.. ct. (\ i1 \ n1. setsum (\ i2. ct i1 i2) {.. + (\ i2 \ n2. setsum (\ i1. ct i1 i2) {.. ct1 ct2 :: nat \ nat. + setsum ct1 {.. ?phi ct1 ct2 n1 n2" + proof(induct rule: nat_sum_induct[of +"\ n1 n2. \ ct1 ct2 :: nat \ nat. + setsum ct1 {.. ?phi ct1 ct2 n1 n2"], + clarify) + fix n1 n2 :: nat and ct1 ct2 :: "nat \ nat" + assume IH: "\ m1 m2. m1 + m2 < n1 + n2 \ + \ dt1 dt2 :: nat \ nat. + setsum dt1 {.. ?phi dt1 dt2 m1 m2" + and ss: "setsum ct1 {.. ct2 n2") + case True + def dt2 \ "\ i2. if i2 = n2 then ct2 i2 - ct1 n1 else ct2 i2" + have "setsum ct1 {.. i1. i1 \ m1 \ setsum (\ i2. dt i1 i2) {.. i2. i2 \ n2 \ setsum (\ i1. dt i1 i2) {.. "\ i1. if i1 = n1 then ct1 i1 - ct2 n2 else ct1 i1" + have "setsum dt1 {.. i1. i1 \ n1 \ setsum (\ i2. dt i1 i2) {.. i2. i2 \ m2 \ setsum (\ i1. dt i1 i2) {.. + \ b1 b1' b2 b2'. {b1,b1'} \ B1 \ {b2,b2'} \ B2 \ u b1 b2 = u b1' b2' + \ b1 = b1' \ b2 = b2'" + +lemma matrix_count_finite: +assumes B1: "B1 \ {}" "finite B1" and B2: "B2 \ {}" "finite B2" and u: "inj2 u B1 B2" +and ss: "setsum N1 B1 = setsum N2 B2" +shows "\ M :: 'a \ nat. + (\ b1 \ B1. setsum (\ b2. M (u b1 b2)) B2 = N1 b1) \ + (\ b2 \ B2. setsum (\ b1. M (u b1 b2)) B1 = N2 b2)" +proof- + obtain n1 where "card B1 = Suc n1" using B1 by (metis card_insert finite.simps) + then obtain e1 where e1: "bij_betw e1 {.. "inv_into {.. i1. i1 < Suc n1 \ f1 (e1 i1) = i1" + and e1f1[simp]: "\ b1. b1 \ B1 \ e1 (f1 b1) = b1" unfolding f1_def + apply (metis bij_betw_inv_into e1, metis bij_betw_inv_into_left e1 lessThan_iff) + by (metis e1_surj f_inv_into_f) + (* *) + obtain n2 where "card B2 = Suc n2" using B2 by (metis card_insert finite.simps) + then obtain e2 where e2: "bij_betw e2 {.. "inv_into {.. i2. i2 < Suc n2 \ f2 (e2 i2) = i2" + and e2f2[simp]: "\ b2. b2 \ B2 \ e2 (f2 b2) = b2" unfolding f2_def + apply (metis bij_betw_inv_into e2, metis bij_betw_inv_into_left e2 lessThan_iff) + by (metis e2_surj f_inv_into_f) + (* *) + let ?ct1 = "N1 o e1" let ?ct2 = "N2 o e2" + have ss: "setsum ?ct1 {.. i1. i1 \ n1 \ setsum (\ i2. ct i1 i2) {.. i2. i2 \ n2 \ setsum (\ i1. ct i1 i2) {.. "{u b1 b2 | b1 b2. b1 \ B1 \ b2 \ B2}" + have "\ a \ A. \ b1b2 \ B1 <*> B2. u (fst b1b2) (snd b1b2) = a" + unfolding A_def Ball_def mem_Collect_eq by auto + then obtain h1h2 where h12: + "\a. a \ A \ u (fst (h1h2 a)) (snd (h1h2 a)) = a \ h1h2 a \ B1 <*> B2" by metis + def h1 \ "fst o h1h2" def h2 \ "snd o h1h2" + have h12[simp]: "\a. a \ A \ u (h1 a) (h2 a) = a" + "\ a. a \ A \ h1 a \ B1" "\ a. a \ A \ h2 a \ B2" + using h12 unfolding h1_def h2_def by force+ + {fix b1 b2 assume b1: "b1 \ B1" and b2: "b2 \ B2" + hence inA: "u b1 b2 \ A" unfolding A_def by auto + hence "u b1 b2 = u (h1 (u b1 b2)) (h2 (u b1 b2))" by auto + moreover have "h1 (u b1 b2) \ B1" "h2 (u b1 b2) \ B2" using inA by auto + ultimately have "h1 (u b1 b2) = b1 \ h2 (u b1 b2) = b2" + using u b1 b2 unfolding inj2_def by fastforce + } + hence h1[simp]: "\ b1 b2. \b1 \ B1; b2 \ B2\ \ h1 (u b1 b2) = b1" and + h2[simp]: "\ b1 b2. \b1 \ B1; b2 \ B2\ \ h2 (u b1 b2) = b2" by auto + def M \ "\ a. ct (f1 (h1 a)) (f2 (h2 a))" + show ?thesis + apply(rule exI[of _ M]) proof safe + fix b1 assume b1: "b1 \ B1" + hence f1b1: "f1 b1 \ n1" using f1 unfolding bij_betw_def + by (metis bij_betwE f1 lessThan_iff less_Suc_eq_le) + have "(\b2\B2. M (u b1 b2)) = (\i2b2\B2. M (u b1 b2)) = N1 b1" . + next + fix b2 assume b2: "b2 \ B2" + hence f2b2: "f2 b2 \ n2" using f2 unfolding bij_betw_def + by (metis bij_betwE f2 lessThan_iff less_Suc_eq_le) + have "(\b1\B1. M (u b1 b2)) = (\i1b1\B1. M (u b1 b2)) = N2 b2" . + qed +qed + +lemma supp_vimage_mmap: +assumes "M \ multiset" +shows "supp M \ f -` (supp (mmap f M))" +using assms by (auto simp: mmap_image) + +lemma mmap_ge_0: +assumes "M \ multiset" +shows "0 < mmap f M b \ (\a. 0 < M a \ f a = b)" +proof- + have f: "finite {a. f a = b \ 0 < M a}" using assms unfolding multiset_def by auto + show ?thesis unfolding mmap_def setsum_gt_0_iff[OF f] by auto +qed + +lemma finite_twosets: +assumes "finite B1" and "finite B2" +shows "finite {u b1 b2 |b1 b2. b1 \ B1 \ b2 \ B2}" (is "finite ?A") +proof- + have A: "?A = (\ b1b2. u (fst b1b2) (snd b1b2)) ` (B1 <*> B2)" by force + show ?thesis unfolding A using finite_cartesian_product[OF assms] by auto +qed + +lemma wp_mmap: +fixes A :: "'a set" and B1 :: "'b1 set" and B2 :: "'b2 set" +assumes wp: "wpull A B1 B2 f1 f2 p1 p2" +shows +"wpull {M. M \ multiset \ supp M \ A} + {N1. N1 \ multiset \ supp N1 \ B1} {N2. N2 \ multiset \ supp N2 \ B2} + (mmap f1) (mmap f2) (mmap p1) (mmap p2)" +unfolding wpull_def proof (safe, unfold Bex_def mem_Collect_eq) + fix N1 :: "'b1 \ nat" and N2 :: "'b2 \ nat" + assume mmap': "mmap f1 N1 = mmap f2 N2" + and N1[simp]: "N1 \ multiset" "supp N1 \ B1" + and N2[simp]: "N2 \ multiset" "supp N2 \ B2" + have mN1[simp]: "mmap f1 N1 \ multiset" using N1 by (auto simp: mmap) + have mN2[simp]: "mmap f2 N2 \ multiset" using N2 by (auto simp: mmap) + def P \ "mmap f1 N1" + have P1: "P = mmap f1 N1" and P2: "P = mmap f2 N2" unfolding P_def using mmap' by auto + note P = P1 P2 + have P_mult[simp]: "P \ multiset" unfolding P_def using N1 by auto + have fin_N1[simp]: "finite (supp N1)" using N1(1) unfolding multiset_def by auto + have fin_N2[simp]: "finite (supp N2)" using N2(1) unfolding multiset_def by auto + have fin_P[simp]: "finite (supp P)" using P_mult unfolding multiset_def by auto + (* *) + def set1 \ "\ c. {b1 \ supp N1. f1 b1 = c}" + have set1[simp]: "\ c b1. b1 \ set1 c \ f1 b1 = c" unfolding set1_def by auto + have fin_set1: "\ c. c \ supp P \ finite (set1 c)" + using N1(1) unfolding set1_def multiset_def by auto + have set1_NE: "\ c. c \ supp P \ set1 c \ {}" + unfolding set1_def P1 mmap_ge_0[OF N1(1)] by auto + have supp_N1_set1: "supp N1 = (\ c \ supp P. set1 c)" + using supp_vimage_mmap[OF N1(1), of f1] unfolding set1_def P1 by auto + hence set1_inclN1: "\c. c \ supp P \ set1 c \ supp N1" by auto + hence set1_incl: "\ c. c \ supp P \ set1 c \ B1" using N1(2) by blast + have set1_disj: "\ c c'. c \ c' \ set1 c \ set1 c' = {}" + unfolding set1_def by auto + have setsum_set1: "\ c. setsum N1 (set1 c) = P c" + unfolding P1 set1_def mmap_def apply(rule setsum_cong) by auto + (* *) + def set2 \ "\ c. {b2 \ supp N2. f2 b2 = c}" + have set2[simp]: "\ c b2. b2 \ set2 c \ f2 b2 = c" unfolding set2_def by auto + have fin_set2: "\ c. c \ supp P \ finite (set2 c)" + using N2(1) unfolding set2_def multiset_def by auto + have set2_NE: "\ c. c \ supp P \ set2 c \ {}" + unfolding set2_def P2 mmap_ge_0[OF N2(1)] by auto + have supp_N2_set2: "supp N2 = (\ c \ supp P. set2 c)" + using supp_vimage_mmap[OF N2(1), of f2] unfolding set2_def P2 by auto + hence set2_inclN2: "\c. c \ supp P \ set2 c \ supp N2" by auto + hence set2_incl: "\ c. c \ supp P \ set2 c \ B2" using N2(2) by blast + have set2_disj: "\ c c'. c \ c' \ set2 c \ set2 c' = {}" + unfolding set2_def by auto + have setsum_set2: "\ c. setsum N2 (set2 c) = P c" + unfolding P2 set2_def mmap_def apply(rule setsum_cong) by auto + (* *) + have ss: "\ c. c \ supp P \ setsum N1 (set1 c) = setsum N2 (set2 c)" + unfolding setsum_set1 setsum_set2 .. + have "\ c \ supp P. \ b1b2 \ (set1 c) \ (set2 c). + \ a \ A. p1 a = fst b1b2 \ p2 a = snd b1b2" + using wp set1_incl set2_incl unfolding wpull_def Ball_def mem_Collect_eq + by simp (metis set1 set2 set_rev_mp) + then obtain uu where uu: + "\ c \ supp P. \ b1b2 \ (set1 c) \ (set2 c). + uu c b1b2 \ A \ p1 (uu c b1b2) = fst b1b2 \ p2 (uu c b1b2) = snd b1b2" by metis + def u \ "\ c b1 b2. uu c (b1,b2)" + have u[simp]: + "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ u c b1 b2 \ A" + "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ p1 (u c b1 b2) = b1" + "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ p2 (u c b1 b2) = b2" + using uu unfolding u_def by auto + {fix c assume c: "c \ supp P" + have "inj2 (u c) (set1 c) (set2 c)" unfolding inj2_def proof clarify + fix b1 b1' b2 b2' + assume "{b1, b1'} \ set1 c" "{b2, b2'} \ set2 c" and 0: "u c b1 b2 = u c b1' b2'" + hence "p1 (u c b1 b2) = b1 \ p2 (u c b1 b2) = b2 \ + p1 (u c b1' b2') = b1' \ p2 (u c b1' b2') = b2'" + using u(2)[OF c] u(3)[OF c] by simp metis + thus "b1 = b1' \ b2 = b2'" using 0 by auto + qed + } note inj = this + def sset \ "\ c. {u c b1 b2 | b1 b2. b1 \ set1 c \ b2 \ set2 c}" + have fin_sset[simp]: "\ c. c \ supp P \ finite (sset c)" unfolding sset_def + using fin_set1 fin_set2 finite_twosets by blast + have sset_A: "\ c. c \ supp P \ sset c \ A" unfolding sset_def by auto + {fix c a assume c: "c \ supp P" and ac: "a \ sset c" + then obtain b1 b2 where b1: "b1 \ set1 c" and b2: "b2 \ set2 c" + and a: "a = u c b1 b2" unfolding sset_def by auto + have "p1 a \ set1 c" and p2a: "p2 a \ set2 c" + using ac a b1 b2 c u(2) u(3) by simp+ + hence "u c (p1 a) (p2 a) = a" unfolding a using b1 b2 inj[OF c] + unfolding inj2_def by (metis c u(2) u(3)) + } note u_p12[simp] = this + {fix c a assume c: "c \ supp P" and ac: "a \ sset c" + hence "p1 a \ set1 c" unfolding sset_def by auto + }note p1[simp] = this + {fix c a assume c: "c \ supp P" and ac: "a \ sset c" + hence "p2 a \ set2 c" unfolding sset_def by auto + }note p2[simp] = this + (* *) + {fix c assume c: "c \ supp P" + hence "\ M. (\ b1 \ set1 c. setsum (\ b2. M (u c b1 b2)) (set2 c) = N1 b1) \ + (\ b2 \ set2 c. setsum (\ b1. M (u c b1 b2)) (set1 c) = N2 b2)" + unfolding sset_def + using matrix_count_finite[OF set1_NE[OF c] fin_set1[OF c] + set2_NE[OF c] fin_set2[OF c] inj[OF c] ss[OF c]] by auto + } + then obtain Ms where + ss1: "\ c b1. \c \ supp P; b1 \ set1 c\ \ + setsum (\ b2. Ms c (u c b1 b2)) (set2 c) = N1 b1" and + ss2: "\ c b2. \c \ supp P; b2 \ set2 c\ \ + setsum (\ b1. Ms c (u c b1 b2)) (set1 c) = N2 b2" + by metis + def SET \ "\ c \ supp P. sset c" + have fin_SET[simp]: "finite SET" unfolding SET_def apply(rule finite_UN_I) by auto + have SET_A: "SET \ A" unfolding SET_def using sset_A by auto + have u_SET[simp]: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ \ u c b1 b2 \ SET" + unfolding SET_def sset_def by blast + {fix c a assume c: "c \ supp P" and a: "a \ SET" and p1a: "p1 a \ set1 c" + then obtain c' where c': "c' \ supp P" and ac': "a \ sset c'" + unfolding SET_def by auto + hence "p1 a \ set1 c'" unfolding sset_def by auto + hence eq: "c = c'" using p1a c c' set1_disj by auto + hence "a \ sset c" using ac' by simp + } note p1_rev = this + {fix c a assume c: "c \ supp P" and a: "a \ SET" and p2a: "p2 a \ set2 c" + then obtain c' where c': "c' \ supp P" and ac': "a \ sset c'" + unfolding SET_def by auto + hence "p2 a \ set2 c'" unfolding sset_def by auto + hence eq: "c = c'" using p2a c c' set2_disj by auto + hence "a \ sset c" using ac' by simp + } note p2_rev = this + (* *) + have "\ a \ SET. \ c \ supp P. a \ sset c" unfolding SET_def by auto + then obtain h where h: "\ a \ SET. h a \ supp P \ a \ sset (h a)" by metis + have h_u[simp]: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ + \ h (u c b1 b2) = c" + by (metis h p2 set2 u(3) u_SET) + have h_u1: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ + \ h (u c b1 b2) = f1 b1" + using h unfolding sset_def by auto + have h_u2: "\ c b1 b2. \c \ supp P; b1 \ set1 c; b2 \ set2 c\ + \ h (u c b1 b2) = f2 b2" + using h unfolding sset_def by auto + def M \ "\ a. if a \ SET \ p1 a \ supp N1 \ p2 a \ supp N2 then Ms (h a) a else 0" + have sM: "supp M \ SET" "supp M \ p1 -` (supp N1)" "supp M \ p2 -` (supp N2)" + unfolding M_def by auto + show "\M. (M \ multiset \ supp M \ A) \ mmap p1 M = N1 \ mmap p2 M = N2" + proof(rule exI[of _ M], safe) + show "M \ multiset" + unfolding multiset_def using finite_subset[OF sM(1) fin_SET] by simp + next + fix a assume "0 < M a" + thus "a \ A" unfolding M_def using SET_A by (cases "a \ SET") auto + next + show "mmap p1 M = N1" + unfolding mmap_def[abs_def] proof(rule ext) + fix b1 + let ?K = "{a. p1 a = b1 \ 0 < M a}" + show "setsum M ?K = N1 b1" + proof(cases "b1 \ supp N1") + case False + hence "?K = {}" using sM(2) by auto + thus ?thesis using False by auto + next + case True + def c \ "f1 b1" + have c: "c \ supp P" and b1: "b1 \ set1 c" + unfolding set1_def c_def P1 using True by (auto simp: mmap_image) + have "setsum M ?K = setsum M {a. p1 a = b1 \ a \ SET}" + apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto + also have "... = setsum M ((\ b2. u c b1 b2) ` (set2 c))" + apply(rule setsum_cong) using c b1 proof safe + fix a assume p1a: "p1 a \ set1 c" and "0 < P c" and "a \ SET" + hence ac: "a \ sset c" using p1_rev by auto + hence "a = u c (p1 a) (p2 a)" using c by auto + moreover have "p2 a \ set2 c" using ac c by auto + ultimately show "a \ u c (p1 a) ` set2 c" by auto + next + fix b2 assume b1: "b1 \ set1 c" and b2: "b2 \ set2 c" + hence "u c b1 b2 \ SET" using c by auto + qed auto + also have "... = setsum (\ b2. M (u c b1 b2)) (set2 c)" + unfolding comp_def[symmetric] apply(rule setsum_reindex) + using inj unfolding inj_on_def inj2_def using b1 c u(3) by blast + also have "... = N1 b1" unfolding ss1[OF c b1, symmetric] + apply(rule setsum_cong[OF refl]) unfolding M_def + using True h_u[OF c b1] set2_def u(2,3)[OF c b1] u_SET[OF c b1] by fastforce + finally show ?thesis . + qed + qed + next + show "mmap p2 M = N2" + unfolding mmap_def[abs_def] proof(rule ext) + fix b2 + let ?K = "{a. p2 a = b2 \ 0 < M a}" + show "setsum M ?K = N2 b2" + proof(cases "b2 \ supp N2") + case False + hence "?K = {}" using sM(3) by auto + thus ?thesis using False by auto + next + case True + def c \ "f2 b2" + have c: "c \ supp P" and b2: "b2 \ set2 c" + unfolding set2_def c_def P2 using True by (auto simp: mmap_image) + have "setsum M ?K = setsum M {a. p2 a = b2 \ a \ SET}" + apply(rule setsum_mono_zero_cong_left) unfolding M_def by auto + also have "... = setsum M ((\ b1. u c b1 b2) ` (set1 c))" + apply(rule setsum_cong) using c b2 proof safe + fix a assume p2a: "p2 a \ set2 c" and "0 < P c" and "a \ SET" + hence ac: "a \ sset c" using p2_rev by auto + hence "a = u c (p1 a) (p2 a)" using c by auto + moreover have "p1 a \ set1 c" using ac c by auto + ultimately show "a \ (\b1. u c b1 (p2 a)) ` set1 c" by auto + next + fix b2 assume b1: "b1 \ set1 c" and b2: "b2 \ set2 c" + hence "u c b1 b2 \ SET" using c by auto + qed auto + also have "... = setsum (M o (\ b1. u c b1 b2)) (set1 c)" + apply(rule setsum_reindex) + using inj unfolding inj_on_def inj2_def using b2 c u(2) by blast + also have "... = setsum (\ b1. M (u c b1 b2)) (set1 c)" + unfolding comp_def[symmetric] by simp + also have "... = N2 b2" unfolding ss2[OF c b2, symmetric] + apply(rule setsum_cong[OF refl]) unfolding M_def set2_def + using True h_u1[OF c _ b2] u(2,3)[OF c _ b2] u_SET[OF c _ b2] + unfolding set1_def by fastforce + finally show ?thesis . + qed + qed + qed +qed + +definition mset_map :: "('a \ 'b) \ 'a multiset \ 'b multiset" where +"mset_map h = Abs_multiset \ mmap h \ count" + +bnf_def mset = mset_map [set_of] "\_::'a multiset. natLeq" ["{#}"] +unfolding mset_map_def +proof - + show "Abs_multiset \ mmap id \ count = id" unfolding mmap_id by (auto simp: count_inverse) +next + fix f g + show "Abs_multiset \ mmap (g \ f) \ count = + Abs_multiset \ mmap g \ count \ (Abs_multiset \ mmap f \ count)" + unfolding comp_def apply(rule ext) + by (auto simp: Abs_multiset_inverse count mmap_comp1 mmap) +next + fix M f g assume eq: "\a. a \ set_of M \ f a = g a" + thus "(Abs_multiset \ mmap f \ count) M = (Abs_multiset \ mmap g \ count) M" apply auto + unfolding cIm_def[abs_def] image_def + by (auto intro!: mmap_cong simp: Abs_multiset_inject count mmap) +next + fix f show "set_of \ (Abs_multiset \ mmap f \ count) = op ` f \ set_of" + by (auto simp: count mmap mmap_image set_of_Abs_multiset supp_count) +next + show "card_order natLeq" by (rule natLeq_card_order) +next + show "cinfinite natLeq" by (rule natLeq_cinfinite) +next + fix M show "|set_of M| \o natLeq" + apply(rule ordLess_imp_ordLeq) + unfolding finite_iff_ordLess_natLeq[symmetric] using finite_set_of . +next + fix A :: "'a set" + have "|{M. set_of M \ A}| \o |{as. set as \ A}|" using card_of_set_of . + also have "|{as. set as \ A}| \o ( |A| +c ctwo) ^c natLeq" + by (rule list_in_bd) + finally show "|{M. set_of M \ A}| \o ( |A| +c ctwo) ^c natLeq" . +next + fix A B1 B2 f1 f2 p1 p2 + let ?map = "\ f. Abs_multiset \ mmap f \ count" + assume wp: "wpull A B1 B2 f1 f2 p1 p2" + show "wpull {x. set_of x \ A} {x. set_of x \ B1} {x. set_of x \ B2} + (?map f1) (?map f2) (?map p1) (?map p2)" + unfolding wpull_def proof safe + fix y1 y2 + assume y1: "set_of y1 \ B1" and y2: "set_of y2 \ B2" + and m: "?map f1 y1 = ?map f2 y2" + def N1 \ "count y1" def N2 \ "count y2" + have "N1 \ multiset \ supp N1 \ B1" and "N2 \ multiset \ supp N2 \ B2" + and "mmap f1 N1 = mmap f2 N2" + using y1 y2 m unfolding N1_def N2_def + by (auto simp: Abs_multiset_inject count mmap) + then obtain M where M: "M \ multiset \ supp M \ A" + and N1: "mmap p1 M = N1" and N2: "mmap p2 M = N2" + using wp_mmap[OF wp] unfolding wpull_def by auto + def x \ "Abs_multiset M" + show "\x\{x. set_of x \ A}. ?map p1 x = y1 \ ?map p2 x = y2" + apply(intro bexI[of _ x]) using M N1 N2 unfolding N1_def N2_def x_def + by (auto simp: count_inverse Abs_multiset_inverse) + qed +qed (unfold set_of_empty, auto) + +end