diff --git a/src/HOL/Analysis/Brouwer_Fixpoint.thy b/src/HOL/Analysis/Brouwer_Fixpoint.thy --- a/src/HOL/Analysis/Brouwer_Fixpoint.thy +++ b/src/HOL/Analysis/Brouwer_Fixpoint.thy @@ -1,2524 +1,2524 @@ (* Author: John Harrison Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) and LCP *) (* At the moment this is just Brouwer's fixpoint theorem. The proof is from *) (* Kuhn: "some combinatorial lemmas in topology", IBM J. v4. (1960) p. 518 *) (* See "http://www.research.ibm.com/journal/rd/045/ibmrd0405K.pdf". *) (* *) (* The script below is quite messy, but at least we avoid formalizing any *) (* topological machinery; we don't even use barycentric subdivision; this is *) (* the big advantage of Kuhn's proof over the usual Sperner's lemma one. *) (* *) (* (c) Copyright, John Harrison 1998-2008 *) section \Brouwer's Fixed Point Theorem\ theory Brouwer_Fixpoint imports Homeomorphism Derivative begin subsection \Retractions\ lemma retract_of_contractible: assumes "contractible T" "S retract_of T" shows "contractible S" using assms apply (clarsimp simp add: retract_of_def contractible_def retraction_def homotopic_with image_subset_iff_funcset) apply (rule_tac x="r a" in exI) apply (rule_tac x="r \ h" in exI) apply (intro conjI continuous_intros continuous_on_compose) apply (erule continuous_on_subset | force)+ done lemma retract_of_path_connected: "\path_connected T; S retract_of T\ \ path_connected S" by (metis path_connected_continuous_image retract_of_def retraction) lemma retract_of_simply_connected: - "\simply_connected T; S retract_of T\ \ simply_connected S" - apply (simp add: retract_of_def retraction_def Pi_iff, clarify) - apply (rule simply_connected_retraction_gen) - apply (force elim!: continuous_on_subset)+ - done + assumes T: "simply_connected T" and "S retract_of T" + shows "simply_connected S" +proof - + obtain r where r: "retraction T S r" + using assms by (metis retract_of_def) + have "S \ T" + by (meson \retraction T S r\ retraction) + then have "(\a. a) \ S \ T" + by blast + then show ?thesis + using simply_connected_retraction_gen [OF T] + by (metis (no_types) r retraction retraction_refl) +qed lemma retract_of_homotopically_trivial: assumes ts: "T retract_of S" and hom: "\f g. \continuous_on U f; f \ U \ S; continuous_on U g; g \ U \ S\ \ homotopic_with_canon (\x. True) U S f g" and "continuous_on U f" "f \ U \ T" and "continuous_on U g" "g \ U \ T" shows "homotopic_with_canon (\x. True) U T f g" proof - obtain r where "r \ S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" using ts by (auto simp: retract_of_def retraction) then obtain k where "Retracts S r T k" unfolding Retracts_def using continuous_on_id by blast then show ?thesis - apply (rule Retracts.homotopically_trivial_retraction_gen) - using assms - apply (force simp: hom)+ - done + by (rule Retracts.homotopically_trivial_retraction_gen) (use assms hom in force)+ qed lemma retract_of_homotopically_trivial_null: assumes ts: "T retract_of S" and hom: "\f. \continuous_on U f; f \ U \ S\ \ \c. homotopic_with_canon (\x. True) U S f (\x. c)" and "continuous_on U f" "f \ U \ T" obtains c where "homotopic_with_canon (\x. True) U T f (\x. c)" proof - obtain r where "r \ S \ S" "continuous_on S r" "\x\S. r (r x) = r x" "T = r ` S" using ts by (auto simp: retract_of_def retraction) then obtain k where "Retracts S r T k" unfolding Retracts_def by fastforce then show ?thesis - apply (rule Retracts.homotopically_trivial_retraction_null_gen) - apply (rule TrueI refl assms that | assumption)+ - done + proof (rule Retracts.homotopically_trivial_retraction_null_gen) + show "\f. \continuous_on U f; f \ U \ S\ + \ \c. homotopic_with_canon (\a. True) U S f (\x. c)" + using hom by blast + qed (use assms that in auto) qed lemma retraction_openin_vimage_iff: "openin (top_of_set S) (S \ r -` U) \ openin (top_of_set T) U" if "retraction S T r" and "U \ T" by (simp add: retraction_openin_vimage_iff that) lemma retract_of_locally_compact: fixes S :: "'a :: {heine_borel,real_normed_vector} set" shows "\ locally compact S; T retract_of S\ \ locally compact T" by (metis locally_compact_closedin closedin_retract) lemma homotopic_into_retract: "\f \ S \ T; g \ S \ T; T retract_of U; homotopic_with_canon (\x. True) S U f g\ \ homotopic_with_canon (\x. True) S T f g" apply (subst (asm) homotopic_with_def) apply (simp add: homotopic_with retract_of_def retraction_def Pi_iff, clarify) apply (rule_tac x="r \ h" in exI) by (smt (verit, ccfv_SIG) comp_def continuous_on_compose continuous_on_subset image_subset_iff) lemma retract_of_locally_connected: assumes "locally connected T" "S retract_of T" shows "locally connected S" using assms - by (metis Abstract_Topology_2.retraction_openin_vimage_iff idempotent_imp_retraction locally_connected_quotient_image retract_ofE) + by (metis retraction_openin_vimage_iff idempotent_imp_retraction locally_connected_quotient_image retract_ofE) lemma retract_of_locally_path_connected: assumes "locally path_connected T" "S retract_of T" shows "locally path_connected S" using assms - by (metis Abstract_Topology_2.retraction_openin_vimage_iff idempotent_imp_retraction locally_path_connected_quotient_image retract_ofE) + by (metis retraction_openin_vimage_iff idempotent_imp_retraction locally_path_connected_quotient_image retract_ofE) text \A few simple lemmas about deformation retracts\ lemma deformation_retract_imp_homotopy_eqv: fixes S :: "'a::euclidean_space set" assumes "homotopic_with_canon (\x. True) S S id r" and r: "retraction S T r" shows "S homotopy_eqv T" proof - have "homotopic_with_canon (\x. True) S S (id \ r) id" by (simp add: assms(1) homotopic_with_symD) moreover have "homotopic_with_canon (\x. True) T T (r \ id) id" using r unfolding retraction_def by (metis eq_id_iff homotopic_with_id2 topspace_euclidean_subtopology) ultimately show ?thesis unfolding homotopy_equivalent_space_def by (smt (verit, del_insts) continuous_map_id continuous_map_subtopology_eu id_def r retraction retraction_comp subset_refl) qed lemma deformation_retract: fixes S :: "'a::euclidean_space set" shows "(\r. homotopic_with_canon (\x. True) S S id r \ retraction S T r) \ T retract_of S \ (\f. homotopic_with_canon (\x. True) S S id f \ f \ S \ T)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (auto simp: retract_of_def retraction_def) next - assume ?rhs - then show ?lhs - apply (clarsimp simp add: retract_of_def retraction_def) - apply (rule_tac x=r in exI, simp) - apply (rule homotopic_with_trans, assumption) - apply (rule_tac f = "r \ f" and g="r \ id" in homotopic_with_eq) - apply (rule_tac Y=S in homotopic_with_compose_continuous_left) + assume R: ?rhs + have "\r f. \T \ S; continuous_on S r; homotopic_with_canon (\x. True) S S id f; + f \ S \ T; r \ S \ T; \x\T. r x = x\ + \ homotopic_with_canon (\x. True) S S f r" + apply (rule_tac f = "r \ f" and g="r \ id" in homotopic_with_eq) + apply (rule_tac Y=S in homotopic_with_compose_continuous_left) apply (auto simp: homotopic_with_sym Pi_iff) done + with R homotopic_with_trans show ?lhs + unfolding retract_of_def retraction_def by blast qed lemma deformation_retract_of_contractible_sing: fixes S :: "'a::euclidean_space set" assumes "contractible S" "a \ S" obtains r where "homotopic_with_canon (\x. True) S S id r" "retraction S {a} r" proof - have "{a} retract_of S" by (simp add: \a \ S\) moreover have "homotopic_with_canon (\x. True) S S id (\x. a)" using assms by (auto simp: contractible_def homotopic_into_contractible image_subset_iff) moreover have "(\x. a) \ S \ {a}" by (simp add: image_subsetI) ultimately show ?thesis by (metis that deformation_retract) qed lemma continuous_on_compact_surface_projection_aux: fixes S :: "'a::t2_space set" assumes "compact S" "S \ T" "image q T \ S" and contp: "continuous_on T p" and "\x. x \ S \ q x = x" and [simp]: "\x. x \ T \ q(p x) = q x" and "\x. x \ T \ p(q x) = p x" shows "continuous_on T q" proof - have *: "image p T = image p S" using assms by auto (metis imageI subset_iff) have contp': "continuous_on S p" by (rule continuous_on_subset [OF contp \S \ T\]) have "continuous_on (p ` T) q" by (simp add: "*" assms(1) assms(2) assms(5) continuous_on_inv contp' rev_subsetD) then have "continuous_on T (q \ p)" by (rule continuous_on_compose [OF contp]) then show ?thesis by (rule continuous_on_eq [of _ "q \ p"]) (simp add: o_def) qed lemma continuous_on_compact_surface_projection: fixes S :: "'a::real_normed_vector set" assumes "compact S" and S: "S \ V - {0}" and "cone V" and iff: "\x k. x \ V - {0} \ 0 < k \ (k *\<^sub>R x) \ S \ d x = k" shows "continuous_on (V - {0}) (\x. d x *\<^sub>R x)" proof (rule continuous_on_compact_surface_projection_aux [OF \compact S\ S]) show "(\x. d x *\<^sub>R x) ` (V - {0}) \ S" using iff by auto show "continuous_on (V - {0}) (\x. inverse(norm x) *\<^sub>R x)" by (intro continuous_intros) force show "\x. x \ S \ d x *\<^sub>R x = x" by (metis S zero_less_one local.iff scaleR_one subset_eq) show "d (x /\<^sub>R norm x) *\<^sub>R (x /\<^sub>R norm x) = d x *\<^sub>R x" if "x \ V - {0}" for x using iff [of "inverse(norm x) *\<^sub>R x" "norm x * d x", symmetric] iff that \cone V\ by (simp add: field_simps cone_def zero_less_mult_iff) show "d x *\<^sub>R x /\<^sub>R norm (d x *\<^sub>R x) = x /\<^sub>R norm x" if "x \ V - {0}" for x proof - have "0 < d x" using local.iff that by blast then show ?thesis by simp qed qed subsection \Kuhn Simplices\ lemma bij_betw_singleton_eq: assumes f: "bij_betw f A B" and g: "bij_betw g A B" and a: "a \ A" assumes eq: "(\x. x \ A \ x \ a \ f x = g x)" shows "f a = g a" proof - have "f ` (A - {a}) = g ` (A - {a})" by (intro image_cong) (simp_all add: eq) then have "B - {f a} = B - {g a}" using f g a by (auto simp: bij_betw_def inj_on_image_set_diff set_eq_iff) moreover have "f a \ B" "g a \ B" using f g a by (auto simp: bij_betw_def) ultimately show ?thesis by auto qed lemmas swap_apply1 = swap_apply(1) lemmas swap_apply2 = swap_apply(2) lemma pointwise_minimal_pointwise_maximal: fixes s :: "(nat \ nat) set" assumes "finite s" and "s \ {}" and "\x\s. \y\s. x \ y \ y \ x" shows "\a\s. \x\s. a \ x" and "\a\s. \x\s. x \ a" using assms proof (induct s rule: finite_ne_induct) case (insert b s) assume *: "\x\insert b s. \y\insert b s. x \ y \ y \ x" then obtain u l where "l \ s" "\b\s. l \ b" "u \ s" "\b\s. b \ u" using insert by auto with * show "\a\insert b s. \x\insert b s. a \ x" "\a\insert b s. \x\insert b s. x \ a" by (metis insert_iff order.trans)+ qed auto lemma kuhn_labelling_lemma: fixes P Q :: "'a::euclidean_space \ bool" assumes "\x. P x \ P (f x)" and "\x. P x \ (\i\Basis. Q i \ 0 \ x\i \ x\i \ 1)" shows "\l. (\x.\i\Basis. l x i \ (1::nat)) \ (\x.\i\Basis. P x \ Q i \ (x\i = 0) \ (l x i = 0)) \ (\x.\i\Basis. P x \ Q i \ (x\i = 1) \ (l x i = 1)) \ (\x.\i\Basis. P x \ Q i \ (l x i = 0) \ x\i \ f x\i) \ (\x.\i\Basis. P x \ Q i \ (l x i = 1) \ f x\i \ x\i)" proof - { fix x i let ?R = "\y. (P x \ Q i \ x \ i = 0 \ y = (0::nat)) \ (P x \ Q i \ x \ i = 1 \ y = 1) \ (P x \ Q i \ y = 0 \ x \ i \ f x \ i) \ (P x \ Q i \ y = 1 \ f x \ i \ x \ i)" { assume "P x" "Q i" "i \ Basis" with assms have "0 \ f x \ i \ f x \ i \ 1" by auto } then have "i \ Basis \ ?R 0 \ ?R 1" by auto } then show ?thesis unfolding all_conj_distrib[symmetric] Ball_def (* FIXME: shouldn't this work by metis? *) by (subst choice_iff[symmetric])+ blast qed subsubsection \The key "counting" observation, somewhat abstracted\ lemma kuhn_counting_lemma: fixes bnd compo compo' face S F defines "nF s == card {f\F. face f s \ compo' f}" assumes [simp, intro]: "finite F" \ \faces\ and [simp, intro]: "finite S" \ \simplices\ and "\f. f \ F \ bnd f \ card {s\S. face f s} = 1" and "\f. f \ F \ \ bnd f \ card {s\S. face f s} = 2" and "\s. s \ S \ compo s \ nF s = 1" and "\s. s \ S \ \ compo s \ nF s = 0 \ nF s = 2" and "odd (card {f\F. compo' f \ bnd f})" shows "odd (card {s\S. compo s})" proof - have "(\s | s \ S \ \ compo s. nF s) + (\s | s \ S \ compo s. nF s) = (\s\S. nF s)" by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong) also have "\ = (\s\S. card {f \ {f\F. compo' f \ bnd f}. face f s}) + (\s\S. card {f \ {f\F. compo' f \ \ bnd f}. face f s})" unfolding sum.distrib[symmetric] by (subst card_Un_disjoint[symmetric]) (auto simp: nF_def intro!: sum.cong arg_cong[where f=card]) also have "\ = 1 * card {f\F. compo' f \ bnd f} + 2 * card {f\F. compo' f \ \ bnd f}" using assms(4,5) by (fastforce intro!: arg_cong2[where f="(+)"] sum_multicount) finally have "odd ((\s | s \ S \ \ compo s. nF s) + card {s\S. compo s})" using assms(6,8) by simp moreover have "(\s | s \ S \ \ compo s. nF s) = (\s | s \ S \ \ compo s \ nF s = 0. nF s) + (\s | s \ S \ \ compo s \ nF s = 2. nF s)" using assms(7) by (subst sum.union_disjoint[symmetric]) (fastforce intro!: sum.cong)+ ultimately show ?thesis by auto qed subsubsection \The odd/even result for faces of complete vertices, generalized\ lemma kuhn_complete_lemma: assumes [simp]: "finite simplices" and face: "\f s. face f s \ (\a\s. f = s - {a})" and card_s[simp]: "\s. s \ simplices \ card s = n + 2" and rl_bd: "\s. s \ simplices \ rl ` s \ {..Suc n}" and bnd: "\f s. s \ simplices \ face f s \ bnd f \ card {s\simplices. face f s} = 1" and nbnd: "\f s. s \ simplices \ face f s \ \ bnd f \ card {s\simplices. face f s} = 2" and odd_card: "odd (card {f. (\s\simplices. face f s) \ rl ` f = {..n} \ bnd f})" shows "odd (card {s\simplices. (rl ` s = {..Suc n})})" proof (rule kuhn_counting_lemma) have finite_s[simp]: "\s. s \ simplices \ finite s" by (metis add_is_0 zero_neq_numeral card.infinite assms(3)) let ?F = "{f. \s\simplices. face f s}" have F_eq: "?F = (\s\simplices. \a\s. {s - {a}})" by (auto simp: face) show "finite ?F" using \finite simplices\ unfolding F_eq by auto show "card {s \ simplices. face f s} = 1" if "f \ ?F" "bnd f" for f using bnd that by auto show "card {s \ simplices. face f s} = 2" if "f \ ?F" "\ bnd f" for f using nbnd that by auto show "odd (card {f \ {f. \s\simplices. face f s}. rl ` f = {..n} \ bnd f})" using odd_card by simp fix s assume s[simp]: "s \ simplices" let ?S = "{f \ {f. \s\simplices. face f s}. face f s \ rl ` f = {..n}}" have "?S = (\a. s - {a}) ` {a\s. rl ` (s - {a}) = {..n}}" using s by (fastforce simp: face) then have card_S: "card ?S = card {a\s. rl ` (s - {a}) = {..n}}" by (auto intro!: card_image inj_onI) { assume rl: "rl ` s = {..Suc n}" then have inj_rl: "inj_on rl s" by (intro eq_card_imp_inj_on) auto moreover obtain a where "rl a = Suc n" "a \ s" by (metis atMost_iff image_iff le_Suc_eq rl) ultimately have n: "{..n} = rl ` (s - {a})" by (auto simp: inj_on_image_set_diff rl) have "{a\s. rl ` (s - {a}) = {..n}} = {a}" using inj_rl \a \ s\ by (auto simp: n inj_on_image_eq_iff[OF inj_rl]) then show "card ?S = 1" unfolding card_S by simp } { assume rl: "rl ` s \ {..Suc n}" show "card ?S = 0 \ card ?S = 2" proof cases assume *: "{..n} \ rl ` s" with rl rl_bd[OF s] have rl_s: "rl ` s = {..n}" by (auto simp: atMost_Suc subset_insert_iff split: if_split_asm) then have "\ inj_on rl s" by (intro pigeonhole) simp then obtain a b where ab: "a \ s" "b \ s" "rl a = rl b" "a \ b" by (auto simp: inj_on_def) then have eq: "rl ` (s - {a}) = rl ` s" by auto with ab have inj: "inj_on rl (s - {a})" by (intro eq_card_imp_inj_on) (auto simp: rl_s card_Diff_singleton_if) { fix x assume "x \ s" "x \ {a, b}" then have "rl ` s - {rl x} = rl ` ((s - {a}) - {x})" by (auto simp: eq inj_on_image_set_diff[OF inj]) also have "\ = rl ` (s - {x})" using ab \x \ {a, b}\ by auto also assume "\ = rl ` s" finally have False using \x\s\ by auto } moreover { fix x assume "x \ {a, b}" with ab have "x \ s \ rl ` (s - {x}) = rl ` s" by (simp add: set_eq_iff image_iff Bex_def) metis } ultimately have "{a\s. rl ` (s - {a}) = {..n}} = {a, b}" unfolding rl_s[symmetric] by fastforce with \a \ b\ show "card ?S = 0 \ card ?S = 2" unfolding card_S by simp next assume "\ {..n} \ rl ` s" then have "\x. rl ` (s - {x}) \ {..n}" by auto then show "card ?S = 0 \ card ?S = 2" unfolding card_S by simp qed } qed fact locale kuhn_simplex = fixes p n and base upd and S :: "(nat \ nat) set" assumes base: "base \ {..< n} \ {..< p}" assumes base_out: "\i. n \ i \ base i = p" assumes upd: "bij_betw upd {..< n} {..< n}" assumes s_pre: "S = (\i j. if j \ upd`{..< i} then Suc (base j) else base j) ` {.. n}" begin definition "enum i j = (if j \ upd`{..< i} then Suc (base j) else base j)" lemma s_eq: "S = enum ` {.. n}" unfolding s_pre enum_def[abs_def] .. lemma upd_space: "i < n \ upd i < n" using upd by (auto dest!: bij_betwE) lemma s_space: "S \ {..< n} \ {.. p}" proof - { fix i assume "i \ n" then have "enum i \ {..< n} \ {.. p}" proof (induct i) case 0 then show ?case using base by (auto simp: Pi_iff less_imp_le enum_def) next case (Suc i) with base show ?case by (auto simp: Pi_iff Suc_le_eq less_imp_le enum_def intro: upd_space) qed } then show ?thesis by (auto simp: s_eq) qed lemma inj_upd: "inj_on upd {..< n}" using upd by (simp add: bij_betw_def) lemma inj_enum: "inj_on enum {.. n}" proof - { fix x y :: nat assume "x \ y" "x \ n" "y \ n" with upd have "upd ` {..< x} \ upd ` {..< y}" by (subst inj_on_image_eq_iff[where C="{..< n}"]) (auto simp: bij_betw_def) then have "enum x \ enum y" by (auto simp: enum_def fun_eq_iff) } then show ?thesis by (auto simp: inj_on_def) qed lemma enum_0: "enum 0 = base" by (simp add: enum_def[abs_def]) lemma base_in_s: "base \ S" unfolding s_eq by (subst enum_0[symmetric]) auto lemma enum_in: "i \ n \ enum i \ S" unfolding s_eq by auto lemma one_step: assumes a: "a \ S" "j < n" assumes *: "\a'. a' \ S \ a' \ a \ a' j = p'" shows "a j \ p'" proof assume "a j = p'" with * a have "\a'. a' \ S \ a' j = p'" by auto then have "\i. i \ n \ enum i j = p'" unfolding s_eq by auto from this[of 0] this[of n] have "j \ upd ` {..< n}" by (auto simp: enum_def fun_eq_iff split: if_split_asm) with upd \j < n\ show False by (auto simp: bij_betw_def) qed lemma upd_inj: "i < n \ j < n \ upd i = upd j \ i = j" using upd by (auto simp: bij_betw_def inj_on_eq_iff) lemma upd_surj: "upd ` {..< n} = {..< n}" using upd by (auto simp: bij_betw_def) lemma in_upd_image: "A \ {..< n} \ i < n \ upd i \ upd ` A \ i \ A" using inj_on_image_mem_iff[of upd "{..< n}"] upd by (auto simp: bij_betw_def) lemma enum_inj: "i \ n \ j \ n \ enum i = enum j \ i = j" using inj_enum by (auto simp: inj_on_eq_iff) lemma in_enum_image: "A \ {.. n} \ i \ n \ enum i \ enum ` A \ i \ A" using inj_on_image_mem_iff[OF inj_enum] by auto lemma enum_mono: "i \ n \ j \ n \ enum i \ enum j \ i \ j" by (auto simp: enum_def le_fun_def in_upd_image Ball_def[symmetric]) lemma enum_strict_mono: "i \ n \ j \ n \ enum i < enum j \ i < j" using enum_mono[of i j] enum_inj[of i j] by (auto simp: le_less) lemma chain: "a \ S \ b \ S \ a \ b \ b \ a" by (auto simp: s_eq enum_mono) lemma less: "a \ S \ b \ S \ a i < b i \ a < b" using chain[of a b] by (auto simp: less_fun_def le_fun_def not_le[symmetric]) lemma enum_0_bot: "a \ S \ a = enum 0 \ (\a'\S. a \ a')" unfolding s_eq by (auto simp: enum_mono Ball_def) lemma enum_n_top: "a \ S \ a = enum n \ (\a'\S. a' \ a)" unfolding s_eq by (auto simp: enum_mono Ball_def) lemma enum_Suc: "i < n \ enum (Suc i) = (enum i)(upd i := Suc (enum i (upd i)))" by (auto simp: fun_eq_iff enum_def upd_inj) lemma enum_eq_p: "i \ n \ n \ j \ enum i j = p" by (induct i) (auto simp: enum_Suc enum_0 base_out upd_space not_less[symmetric]) lemma out_eq_p: "a \ S \ n \ j \ a j = p" unfolding s_eq by (auto simp: enum_eq_p) lemma s_le_p: "a \ S \ a j \ p" using out_eq_p[of a j] s_space by (cases "j < n") auto lemma le_Suc_base: "a \ S \ a j \ Suc (base j)" unfolding s_eq by (auto simp: enum_def) lemma base_le: "a \ S \ base j \ a j" unfolding s_eq by (auto simp: enum_def) lemma enum_le_p: "i \ n \ j < n \ enum i j \ p" using enum_in[of i] s_space by auto lemma enum_less: "a \ S \ i < n \ enum i < a \ enum (Suc i) \ a" unfolding s_eq by (auto simp: enum_strict_mono enum_mono) lemma ksimplex_0: "n = 0 \ S = {(\x. p)}" using s_eq enum_def base_out by auto lemma replace_0: assumes "j < n" "a \ S" and p: "\x\S - {a}. x j = 0" and "x \ S" shows "x \ a" proof cases assume "x \ a" have "a j \ 0" using assms by (intro one_step[where a=a]) auto with less[OF \x\S\ \a\S\, of j] p[rule_format, of x] \x \ S\ \x \ a\ show ?thesis by auto qed simp lemma replace_1: assumes "j < n" "a \ S" and p: "\x\S - {a}. x j = p" and "x \ S" shows "a \ x" proof cases assume "x \ a" have "a j \ p" using assms by (intro one_step[where a=a]) auto with enum_le_p[of _ j] \j < n\ \a\S\ have "a j < p" by (auto simp: less_le s_eq) with less[OF \a\S\ \x\S\, of j] p[rule_format, of x] \x \ S\ \x \ a\ show ?thesis by auto qed simp end locale kuhn_simplex_pair = s: kuhn_simplex p n b_s u_s s + t: kuhn_simplex p n b_t u_t t for p n b_s u_s s b_t u_t t begin lemma enum_eq: assumes l: "i \ l" "l \ j" and "j + d \ n" assumes eq: "s.enum ` {i .. j} = t.enum ` {i + d .. j + d}" shows "s.enum l = t.enum (l + d)" using l proof (induct l rule: dec_induct) case base then have s: "s.enum i \ t.enum ` {i + d .. j + d}" and t: "t.enum (i + d) \ s.enum ` {i .. j}" using eq by auto from t \i \ j\ \j + d \ n\ have "s.enum i \ t.enum (i + d)" by (auto simp: s.enum_mono) moreover from s \i \ j\ \j + d \ n\ have "t.enum (i + d) \ s.enum i" by (auto simp: t.enum_mono) ultimately show ?case by auto next case (step l) moreover from step.prems \j + d \ n\ have "s.enum l < s.enum (Suc l)" "t.enum (l + d) < t.enum (Suc l + d)" by (simp_all add: s.enum_strict_mono t.enum_strict_mono) moreover have "s.enum (Suc l) \ t.enum ` {i + d .. j + d}" "t.enum (Suc l + d) \ s.enum ` {i .. j}" using step \j + d \ n\ eq by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum (Suc l) = t.enum (Suc (l + d))" using \j + d \ n\ by (intro antisym s.enum_less[THEN iffD1] t.enum_less[THEN iffD1]) (auto intro!: s.enum_in t.enum_in) then show ?case by simp qed lemma ksimplex_eq_bot: assumes a: "a \ s" "\a'. a' \ s \ a \ a'" assumes b: "b \ t" "\b'. b' \ t \ b \ b'" assumes eq: "s - {a} = t - {b}" shows "s = t" proof cases assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp next assume "n \ 0" have "s.enum 0 = (s.enum (Suc 0)) (u_s 0 := s.enum (Suc 0) (u_s 0) - 1)" "t.enum 0 = (t.enum (Suc 0)) (u_t 0 := t.enum (Suc 0) (u_t 0) - 1)" using \n \ 0\ by (simp_all add: s.enum_Suc t.enum_Suc) moreover have e0: "a = s.enum 0" "b = t.enum 0" using a b by (simp_all add: s.enum_0_bot t.enum_0_bot) moreover { fix j assume "0 < j" "j \ n" moreover have "s - {a} = s.enum ` {Suc 0 .. n}" "t - {b} = t.enum ` {Suc 0 .. n}" unfolding s.s_eq t.s_eq e0 by (auto simp: s.enum_inj t.enum_inj) ultimately have "s.enum j = t.enum j" using enum_eq[of "1" j n 0] eq by auto } note enum_eq = this then have "s.enum (Suc 0) = t.enum (Suc 0)" using \n \ 0\ by auto moreover { fix j assume "Suc j < n" with enum_eq[of "Suc j"] enum_eq[of "Suc (Suc j)"] have "u_s (Suc j) = u_t (Suc j)" using s.enum_Suc[of "Suc j"] t.enum_Suc[of "Suc j"] by (auto simp: fun_eq_iff split: if_split_asm) } then have "\j. 0 < j \ j < n \ u_s j = u_t j" by (auto simp: gr0_conv_Suc) with \n \ 0\ have "u_t 0 = u_s 0" by (intro bij_betw_singleton_eq[OF t.upd s.upd, of 0]) auto ultimately have "a = b" by simp with assms show "s = t" by auto qed lemma ksimplex_eq_top: assumes a: "a \ s" "\a'. a' \ s \ a' \ a" assumes b: "b \ t" "\b'. b' \ t \ b' \ b" assumes eq: "s - {a} = t - {b}" shows "s = t" proof (cases n) assume "n = 0" with s.ksimplex_0 t.ksimplex_0 show ?thesis by simp next case (Suc n') have "s.enum n = (s.enum n') (u_s n' := Suc (s.enum n' (u_s n')))" "t.enum n = (t.enum n') (u_t n' := Suc (t.enum n' (u_t n')))" using Suc by (simp_all add: s.enum_Suc t.enum_Suc) moreover have en: "a = s.enum n" "b = t.enum n" using a b by (simp_all add: s.enum_n_top t.enum_n_top) moreover { fix j assume "j < n" moreover have "s - {a} = s.enum ` {0 .. n'}" "t - {b} = t.enum ` {0 .. n'}" unfolding s.s_eq t.s_eq en by (auto simp: s.enum_inj t.enum_inj Suc) ultimately have "s.enum j = t.enum j" using enum_eq[of "0" j n' 0] eq Suc by auto } note enum_eq = this then have "s.enum n' = t.enum n'" using Suc by auto moreover { fix j assume "j < n'" with enum_eq[of j] enum_eq[of "Suc j"] have "u_s j = u_t j" using s.enum_Suc[of j] t.enum_Suc[of j] by (auto simp: Suc fun_eq_iff split: if_split_asm) } then have "\j. j < n' \ u_s j = u_t j" by (auto simp: gr0_conv_Suc) then have "u_t n' = u_s n'" by (intro bij_betw_singleton_eq[OF t.upd s.upd, of n']) (auto simp: Suc) ultimately have "a = b" by simp with assms show "s = t" by auto qed end inductive ksimplex for p n :: nat where ksimplex: "kuhn_simplex p n base upd s \ ksimplex p n s" lemma finite_ksimplexes: "finite {s. ksimplex p n s}" proof (rule finite_subset) { fix a s assume "ksimplex p n s" "a \ s" then obtain b u where "kuhn_simplex p n b u s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p n b u s . from s_space \a \ s\ out_eq_p[OF \a \ s\] have "a \ (\f x. if n \ x then p else f x) ` ({..< n} \\<^sub>E {.. p})" by (auto simp: image_iff subset_eq Pi_iff split: if_split_asm intro!: bexI[of _ "restrict a {..< n}"]) } then show "{s. ksimplex p n s} \ Pow ((\f x. if n \ x then p else f x) ` ({..< n} \\<^sub>E {.. p}))" by auto qed (simp add: finite_PiE) lemma ksimplex_card: assumes "ksimplex p n s" shows "card s = Suc n" using assms proof cases case (ksimplex u b) then interpret kuhn_simplex p n u b s . show ?thesis by (simp add: card_image s_eq inj_enum) qed lemma simplex_top_face: assumes "0 < p" "\x\s'. x n = p" shows "ksimplex p n s' \ (\s a. ksimplex p (Suc n) s \ a \ s \ s' = s - {a})" using assms proof safe fix s a assume "ksimplex p (Suc n) s" and a: "a \ s" and na: "\x\s - {a}. x n = p" then show "ksimplex p n (s - {a})" proof cases case (ksimplex base upd) then interpret kuhn_simplex p "Suc n" base upd "s" . have "a n < p" using one_step[of a n p] na \a\s\ s_space by (auto simp: less_le) then have "a = enum 0" using \a \ s\ na by (subst enum_0_bot) (auto simp: le_less intro!: less[of a _ n]) then have s_eq: "s - {a} = enum ` Suc ` {.. n}" using s_eq by (simp add: atMost_Suc_eq_insert_0 insert_ident in_enum_image subset_eq) then have "enum 1 \ s - {a}" by auto then have "upd 0 = n" using \a n < p\ \a = enum 0\ na[rule_format, of "enum 1"] by (auto simp: fun_eq_iff enum_Suc split: if_split_asm) then have "bij_betw upd (Suc ` {..< n}) {..< n}" using upd by (subst notIn_Un_bij_betw3[where b=0]) (auto simp: lessThan_Suc[symmetric] lessThan_Suc_eq_insert_0) then have "bij_betw (upd\Suc) {..enum 1 \ s - {a}\] \a = enum 0\ by (auto simp: \upd 0 = n\) show ?thesis proof (rule ksimplex.intros, standard) show "bij_betw (upd\Suc) {..< n} {..< n}" by fact show "base(n := p) \ {.. {..i. n\i \ (base(n := p)) i = p" using base base_out by (auto simp: Pi_iff) have "\i. Suc ` {..< i} = {..< Suc i} - {0}" by (auto simp: image_iff Ball_def) arith then have upd_Suc: "\i. i \ n \ (upd\Suc) ` {..< i} = upd ` {..< Suc i} - {n}" using \upd 0 = n\ upd_inj by (auto simp add: image_iff less_Suc_eq_0_disj) have n_in_upd: "\i. n \ upd ` {..< Suc i}" using \upd 0 = n\ by auto define f' where "f' i j = (if j \ (upd\Suc)`{..< i} then Suc ((base(n := p)) j) else (base(n := p)) j)" for i j { fix x i assume i [arith]: "i \ n" with upd_Suc have "(upd \ Suc) ` {..a n < p\ \a = enum 0\ \upd 0 = n\ \a n = p - 1\ have "enum (Suc i) x = f' i x" by (auto simp add: f'_def enum_def) } then show "s - {a} = f' ` {.. n}" unfolding s_eq image_comp by (intro image_cong) auto qed qed next assume "ksimplex p n s'" and *: "\x\s'. x n = p" then show "\s a. ksimplex p (Suc n) s \ a \ s \ s' = s - {a}" proof cases case (ksimplex base upd) then interpret kuhn_simplex p n base upd s' . define b where "b = base (n := p - 1)" define u where "u i = (case i of 0 \ n | Suc i \ upd i)" for i have "ksimplex p (Suc n) (s' \ {b})" proof (rule ksimplex.intros, standard) show "b \ {.. {..0 < p\ unfolding lessThan_Suc b_def by (auto simp: PiE_iff) show "\i. Suc n \ i \ b i = p" using base_out by (auto simp: b_def) have "bij_betw u (Suc ` {..< n} \ {0}) ({.. {u 0})" using upd by (intro notIn_Un_bij_betw) (auto simp: u_def bij_betw_def image_comp comp_def inj_on_def) then show "bij_betw u {.. u`{..< i} then Suc (b j) else b j)" for i j have u_eq: "\i. i \ n \ u ` {..< Suc i} = upd ` {..< i} \ { n }" by (auto simp: u_def image_iff upd_inj Ball_def split: nat.split) arith { fix x have "x \ n \ n \ upd ` {.. {0})" unfolding atMost_Suc_eq_insert_0 by simp also have "\ = (f' \ Suc) ` {.. n} \ {b}" by (auto simp: f'_def) also have "(f' \ Suc) ` {.. n} = s'" using \0 < p\ base_out[of n] unfolding s_eq enum_def[abs_def] f'_def[abs_def] upd_space by (intro image_cong) (simp_all add: u_eq b_def fun_eq_iff n_not_upd) finally show "s' \ {b} = f' ` {.. Suc n}" .. qed moreover have "b \ s'" using * \0 < p\ by (auto simp: b_def) ultimately show ?thesis by auto qed qed lemma ksimplex_replace_0: assumes s: "ksimplex p n s" and a: "a \ s" assumes j: "j < n" and p: "\x\s - {a}. x j = 0" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 1" using s proof cases case (ksimplex b_s u_s) { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t by intro_locales fact+ assume b: "b \ t" "t - {b} = s - {a}" with a j p s.replace_0[of _ a] t.replace_0[of _ b] have "s = t" by (intro ksimplex_eq_top[of a b]) auto } then have "{s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = {s}" using s \a \ s\ by auto then show ?thesis by simp qed lemma ksimplex_replace_1: assumes s: "ksimplex p n s" and a: "a \ s" assumes j: "j < n" and p: "\x\s - {a}. x j = p" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 1" using s proof cases case (ksimplex b_s u_s) { fix t b assume "ksimplex p n t" then obtain b_t u_t where "kuhn_simplex p n b_t u_t t" by (auto elim: ksimplex.cases) interpret kuhn_simplex_pair p n b_s u_s s b_t u_t t by intro_locales fact+ assume b: "b \ t" "t - {b} = s - {a}" with a j p s.replace_1[of _ a] t.replace_1[of _ b] have "s = t" by (intro ksimplex_eq_bot[of a b]) auto } then have "{s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = {s}" using s \a \ s\ by auto then show ?thesis by simp qed lemma ksimplex_replace_2: assumes s: "ksimplex p n s" and "a \ s" and "n \ 0" and lb: "\jx\s - {a}. x j \ 0" and ub: "\jx\s - {a}. x j \ p" shows "card {s'. ksimplex p n s' \ (\b\s'. s' - {b} = s - {a})} = 2" using s proof cases case (ksimplex base upd) then interpret kuhn_simplex p n base upd s . from \a \ s\ obtain i where "i \ n" "a = enum i" unfolding s_eq by auto from \i \ n\ have "i = 0 \ i = n \ (0 < i \ i < n)" by linarith then have "\!s'. s' \ s \ ksimplex p n s' \ (\b\s'. s - {a} = s'- {b})" proof (elim disjE conjE) assume "i = 0" define rot where [abs_def]: "rot i = (if i + 1 = n then 0 else i + 1)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" by (auto simp: bij_betw_def inj_on_def image_iff Ball_def rot_def) arith+ from rot upd have "bij_betw ?upd {.. ?upd`{..< i} then Suc (enum (Suc 0) j) else enum (Suc 0) j)" for i j interpret b: kuhn_simplex p n "enum (Suc 0)" "upd \ rot" "f' ` {.. n}" proof from \a = enum i\ ub \n \ 0\ \i = 0\ obtain i' where "i' \ n" "enum i' \ enum 0" "enum i' (upd 0) \ p" unfolding s_eq by (auto intro: upd_space simp: enum_inj) then have "enum 1 \ enum i'" "enum i' (upd 0) < p" using enum_le_p[of i' "upd 0"] by (auto simp: enum_inj enum_mono upd_space) then have "enum 1 (upd 0) < p" by (auto simp: le_fun_def intro: le_less_trans) then show "enum (Suc 0) \ {.. {..n \ 0\ by (auto simp: enum_0 enum_Suc PiE_iff extensional_def upd_space) { fix i assume "n \ i" then show "enum (Suc 0) i = p" using \n \ 0\ by (auto simp: enum_eq_p) } show "bij_betw ?upd {..n \ 0\ show ?thesis by (simp only: f'_def enum_def fun_eq_iff image_comp [symmetric]) (auto simp add: upd_inj) qed then have "enum ` Suc ` {..< n} = f' ` {..< n}" by (force simp: enum_inj) also have "Suc ` {..< n} = {.. n} - {0}" by (auto simp: image_iff Ball_def) arith also have "{..< n} = {.. n} - {n}" by auto finally have eq: "s - {a} = f' ` {.. n} - {f' n}" unfolding s_eq \a = enum i\ \i = 0\ by (simp add: inj_on_image_set_diff[OF inj_enum] inj_on_image_set_diff[OF inj_f']) have "enum 0 < f' 0" using \n \ 0\ by (simp add: enum_strict_mono f'_eq_enum) also have "\ < f' n" using \n \ 0\ b.enum_strict_mono[of 0 n] unfolding b_enum by simp finally have "a \ f' n" using \a = enum i\ \i = 0\ by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b u where "kuhn_simplex p n b u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b u t . { fix x assume "x \ s" "x \ a" then have "x (upd 0) = enum (Suc 0) (upd 0)" by (auto simp: \a = enum i\ \i = 0\ s_eq enum_def enum_inj) } then have eq_upd0: "\x\t-{c}. x (upd 0) = enum (Suc 0) (upd 0)" unfolding eq_sma[symmetric] by auto then have "c (upd 0) \ enum (Suc 0) (upd 0)" using \n \ 0\ by (intro t.one_step[OF \c\t\ ]) (auto simp: upd_space) then have "c (upd 0) < enum (Suc 0) (upd 0) \ c (upd 0) > enum (Suc 0) (upd 0)" by auto then have "t = s \ t = f' ` {..n}" proof (elim disjE conjE) assume *: "c (upd 0) < enum (Suc 0) (upd 0)" interpret st: kuhn_simplex_pair p n base upd s b u t .. { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "c \ x" by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } note top = this have "s = t" using \a = enum i\ \i = 0\ \c \ t\ by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq_sma]) (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) then show ?thesis by simp next assume *: "c (upd 0) > enum (Suc 0) (upd 0)" interpret st: kuhn_simplex_pair p n "enum (Suc 0)" "upd \ rot" "f' ` {.. n}" b u t .. have eq: "f' ` {..n} - {f' n} = t - {c}" using eq_sma eq by simp { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "x \ c" by (auto simp: le_less intro!: t.less[of _ _ "upd 0"]) } note top = this have "f' ` {..n} = t" using \a = enum i\ \i = 0\ \c \ t\ by (intro st.ksimplex_eq_top[OF _ _ _ _ eq]) (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono b_enum[symmetric] top) then show ?thesis by simp qed } with ks_f' eq \a \ f' n\ \n \ 0\ show ?thesis apply (intro ex1I[of _ "f' ` {.. n}"]) - apply auto [] + apply auto [] apply metis done next assume "i = n" from \n \ 0\ obtain n' where n': "n = Suc n'" by (cases n) auto define rot where "rot i = (case i of 0 \ n' | Suc i \ i)" for i let ?upd = "upd \ rot" have rot: "bij_betw rot {..< n} {..< n}" by (auto simp: bij_betw_def inj_on_def image_iff Bex_def rot_def n' split: nat.splits) arith from rot upd have "bij_betw ?upd {.. ?upd`{..< i} then Suc (b j) else b j)" for i j interpret b: kuhn_simplex p n b "upd \ rot" "f' ` {.. n}" proof { fix i assume "n \ i" then show "b i = p" using base_out[of i] upd_space[of n'] by (auto simp: b_def n') } show "b \ {.. {..n \ 0\ upd_space[of n'] by (auto simp: b_def PiE_def Pi_iff Ball_def upd_space extensional_def n') show "bij_betw ?upd {..n \ 0\ by auto { from \a = enum i\ \n \ 0\ \i = n\ lb upd_space[of n'] obtain i' where "i' \ n" "enum i' \ enum n" "0 < enum i' (upd n')" unfolding s_eq by (auto simp: enum_inj n') moreover have "enum i' (upd n') = base (upd n')" unfolding enum_def using \i' \ n\ \enum i' \ enum n\ by (auto simp: n' upd_inj enum_inj) ultimately have "0 < base (upd n')" by auto } then have benum1: "b.enum (Suc 0) = base" unfolding b.enum_Suc[OF \0] b.enum_0 by (auto simp: b_def rot_def) have [simp]: "\j. Suc j < n \ rot ` {..< Suc j} = {n'} \ {..< j}" by (auto simp: rot_def image_iff Ball_def split: nat.splits) have rot_simps: "\j. rot (Suc j) = j" "rot 0 = n'" by (simp_all add: rot_def) { fix j assume j: "Suc j \ n" then have "b.enum (Suc j) = enum j" by (induct j) (auto simp: benum1 enum_0 b.enum_Suc enum_Suc rot_simps) } note b_enum_eq_enum = this then have "enum ` {..< n} = b.enum ` Suc ` {..< n}" by (auto simp: image_comp intro!: image_cong) also have "Suc ` {..< n} = {.. n} - {0}" by (auto simp: image_iff Ball_def) arith also have "{..< n} = {.. n} - {n}" by auto finally have eq: "s - {a} = b.enum ` {.. n} - {b.enum 0}" unfolding s_eq \a = enum i\ \i = n\ using inj_on_image_set_diff[OF inj_enum Diff_subset, of "{n}"] inj_on_image_set_diff[OF b.inj_enum Diff_subset, of "{0}"] by (simp add: comp_def) have "b.enum 0 \ b.enum n" by (simp add: b.enum_mono) also have "b.enum n < enum n" using \n \ 0\ by (simp add: enum_strict_mono b_enum_eq_enum n') finally have "a \ b.enum 0" using \a = enum i\ \i = n\ by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b' u where "kuhn_simplex p n b' u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b' u t . { fix x assume "x \ s" "x \ a" then have "x (upd n') = enum n' (upd n')" by (auto simp: \a = enum i\ n' \i = n\ s_eq enum_def enum_inj in_upd_image) } then have eq_upd0: "\x\t-{c}. x (upd n') = enum n' (upd n')" unfolding eq_sma[symmetric] by auto then have "c (upd n') \ enum n' (upd n')" using \n \ 0\ by (intro t.one_step[OF \c\t\ ]) (auto simp: n' upd_space[unfolded n']) then have "c (upd n') < enum n' (upd n') \ c (upd n') > enum n' (upd n')" by auto then have "t = s \ t = b.enum ` {..n}" proof (elim disjE conjE) assume *: "c (upd n') > enum n' (upd n')" interpret st: kuhn_simplex_pair p n base upd s b' u t .. { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "x \ c" by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } note top = this have "s = t" using \a = enum i\ \i = n\ \c \ t\ by (intro st.ksimplex_eq_top[OF _ _ _ _ eq_sma]) (auto simp: s_eq enum_mono t.s_eq t.enum_mono top) then show ?thesis by simp next assume *: "c (upd n') < enum n' (upd n')" interpret st: kuhn_simplex_pair p n b "upd \ rot" "f' ` {.. n}" b' u t .. have eq: "f' ` {..n} - {b.enum 0} = t - {c}" using eq_sma eq f' by simp { fix x assume "x \ t" with * \c\t\ eq_upd0[rule_format, of x] have "c \ x" by (auto simp: le_less intro!: t.less[of _ _ "upd n'"]) } note bot = this have "f' ` {..n} = t" using \a = enum i\ \i = n\ \c \ t\ by (intro st.ksimplex_eq_bot[OF _ _ _ _ eq]) (auto simp: b.s_eq b.enum_mono t.s_eq t.enum_mono bot) with f' show ?thesis by simp qed } with ks_f' eq \a \ b.enum 0\ \n \ 0\ show ?thesis apply (intro ex1I[of _ "b.enum ` {.. n}"]) apply fastforce apply metis done next assume i: "0 < i" "i < n" define i' where "i' = i - 1" with i have "Suc i' < n" by simp with i have Suc_i': "Suc i' = i" by (simp add: i'_def) let ?upd = "Fun.swap i' i upd" from i upd have "bij_betw ?upd {..< n} {..< n}" by (subst bij_betw_swap_iff) (auto simp: i'_def) define f' where [abs_def]: "f' i j = (if j \ ?upd`{..< i} then Suc (base j) else base j)" for i j interpret b: kuhn_simplex p n base ?upd "f' ` {.. n}" proof show "base \ {.. {.. i" then show "base i = p" by (rule base_out) } show "bij_betw ?upd {.. {..n}" using i by auto { fix j assume "j \ n" with i Suc_i' have "enum j = b.enum j \ j \ i" unfolding fun_eq_iff enum_def b.enum_def image_comp [symmetric] apply (cases \i = j\) apply (metis imageI in_upd_image lessI lessThan_iff lessThan_subset_iff order_less_le transpose_apply_first) by (metis lessThan_iff linorder_not_less not_less_eq_eq order_less_le transpose_image_eq) } note enum_eq_benum = this then have "enum ` ({.. n} - {i}) = b.enum ` ({.. n} - {i})" by (intro image_cong) auto then have eq: "s - {a} = b.enum ` {.. n} - {b.enum i}" unfolding s_eq \a = enum i\ using inj_on_image_set_diff[OF inj_enum Diff_subset \{i} \ {..n}\] inj_on_image_set_diff[OF b.inj_enum Diff_subset \{i} \ {..n}\] by (simp add: comp_def) have "a \ b.enum i" using \a = enum i\ enum_eq_benum i by auto { fix t c assume "ksimplex p n t" "c \ t" and eq_sma: "s - {a} = t - {c}" obtain b' u where "kuhn_simplex p n b' u t" using \ksimplex p n t\ by (auto elim: ksimplex.cases) then interpret t: kuhn_simplex p n b' u t . have "enum i' \ s - {a}" "enum (i + 1) \ s - {a}" using \a = enum i\ i enum_in by (auto simp: enum_inj i'_def) then obtain l k where l: "t.enum l = enum i'" "l \ n" "t.enum l \ c" and k: "t.enum k = enum (i + 1)" "k \ n" "t.enum k \ c" unfolding eq_sma by (auto simp: t.s_eq) with i have "t.enum l < t.enum k" by (simp add: enum_strict_mono i'_def) with \l \ n\ \k \ n\ have "l < k" by (simp add: t.enum_strict_mono) { assume "Suc l = k" have "enum (Suc (Suc i')) = t.enum (Suc l)" using i by (simp add: k \Suc l = k\ i'_def) then have False using \l < k\ \k \ n\ \Suc i' < n\ by (auto simp: t.enum_Suc enum_Suc l upd_inj fun_eq_iff split: if_split_asm) (metis Suc_lessD n_not_Suc_n upd_inj) } with \l < k\ have "Suc l < k" by arith have c_eq: "c = t.enum (Suc l)" proof (rule ccontr) assume "c \ t.enum (Suc l)" then have "t.enum (Suc l) \ s - {a}" using \l < k\ \k \ n\ by (simp add: t.s_eq eq_sma) then obtain j where "t.enum (Suc l) = enum j" "j \ n" "enum j \ enum i" unfolding s_eq \a = enum i\ by auto with i have "t.enum (Suc l) \ t.enum l \ t.enum k \ t.enum (Suc l)" by (auto simp: i'_def enum_mono enum_inj l k) with \Suc l < k\ \k \ n\ show False by (simp add: t.enum_mono) qed { have "t.enum (Suc (Suc l)) \ s - {a}" unfolding eq_sma c_eq t.s_eq using \Suc l < k\ \k \ n\ by (auto simp: t.enum_inj) then obtain j where eq: "t.enum (Suc (Suc l)) = enum j" and "j \ n" "j \ i" by (auto simp: s_eq \a = enum i\) moreover have "enum i' < t.enum (Suc (Suc l))" unfolding l(1)[symmetric] using \Suc l < k\ \k \ n\ by (auto simp: t.enum_strict_mono) ultimately have "i' < j" using i by (simp add: enum_strict_mono i'_def) with \j \ i\ \j \ n\ have "t.enum k \ t.enum (Suc (Suc l))" unfolding i'_def by (simp add: enum_mono k eq) then have "k \ Suc (Suc l)" using \k \ n\ \Suc l < k\ by (simp add: t.enum_mono) } with \Suc l < k\ have "Suc (Suc l) = k" by simp then have "enum (Suc (Suc i')) = t.enum (Suc (Suc l))" using i by (simp add: k i'_def) also have "\ = (enum i') (u l := Suc (enum i' (u l)), u (Suc l) := Suc (enum i' (u (Suc l))))" using \Suc l < k\ \k \ n\ by (simp add: t.enum_Suc l t.upd_inj) finally have "(u l = upd i' \ u (Suc l) = upd (Suc i')) \ (u l = upd (Suc i') \ u (Suc l) = upd i')" using \Suc i' < n\ by (auto simp: enum_Suc fun_eq_iff split: if_split_asm) then have "t = s \ t = b.enum ` {..n}" proof (elim disjE conjE) assume u: "u l = upd i'" have "c = t.enum (Suc l)" unfolding c_eq .. also have "t.enum (Suc l) = enum (Suc i')" using u \l < k\ \k \ n\ \Suc i' < n\ by (simp add: enum_Suc t.enum_Suc l) also have "\ = a" using \a = enum i\ i by (simp add: i'_def) finally show ?thesis using eq_sma \a \ s\ \c \ t\ by auto next assume u: "u l = upd (Suc i')" define B where "B = b.enum ` {..n}" have "b.enum i' = enum i'" using enum_eq_benum[of i'] i by (auto simp: i'_def gr0_conv_Suc) have "c = t.enum (Suc l)" unfolding c_eq .. also have "t.enum (Suc l) = b.enum (Suc i')" using u \l < k\ \k \ n\ \Suc i' < n\ by (simp_all add: enum_Suc t.enum_Suc l b.enum_Suc \b.enum i' = enum i'\) (simp add: Suc_i') also have "\ = b.enum i" using i by (simp add: i'_def) finally have "c = b.enum i" . then have "t - {c} = B - {c}" "c \ B" unfolding eq_sma[symmetric] eq B_def using i by auto with \c \ t\ have "t = B" by auto then show ?thesis by (simp add: B_def) qed } with ks_f' eq \a \ b.enum i\ \n \ 0\ \i \ n\ show ?thesis apply (intro ex1I[of _ "b.enum ` {.. n}"]) apply auto [] apply metis done qed then show ?thesis using s \a \ s\ by (simp add: card_2_iff' Ex1_def) metis qed text \Hence another step towards concreteness.\ lemma kuhn_simplex_lemma: assumes "\s. ksimplex p (Suc n) s \ rl ` s \ {.. Suc n}" and "odd (card {f. \s a. ksimplex p (Suc n) s \ a \ s \ (f = s - {a}) \ rl ` f = {..n} \ ((\j\n. \x\f. x j = 0) \ (\j\n. \x\f. x j = p))})" shows "odd (card {s. ksimplex p (Suc n) s \ rl ` s = {..Suc n}})" proof (rule kuhn_complete_lemma[OF finite_ksimplexes refl, unfolded mem_Collect_eq, where bnd="\f. (\j\{..n}. \x\f. x j = 0) \ (\j\{..n}. \x\f. x j = p)"], safe del: notI) have *: "\x y. x = y \ odd (card x) \ odd (card y)" by auto show "odd (card {f. (\s\{s. ksimplex p (Suc n) s}. \a\s. f = s - {a}) \ rl ` f = {..n} \ ((\j\{..n}. \x\f. x j = 0) \ (\j\{..n}. \x\f. x j = p))})" apply (rule *[OF _ assms(2)]) apply (auto simp: atLeast0AtMost) done next fix s assume s: "ksimplex p (Suc n) s" then show "card s = n + 2" by (simp add: ksimplex_card) fix a assume a: "a \ s" then show "rl a \ Suc n" using assms(1) s by (auto simp: subset_eq) let ?S = "{t. ksimplex p (Suc n) t \ (\b\t. s - {a} = t - {b})}" { fix j assume j: "j \ n" "\x\s - {a}. x j = 0" with s a show "card ?S = 1" using ksimplex_replace_0[of p "n + 1" s a j] by (subst eq_commute) simp } { fix j assume j: "j \ n" "\x\s - {a}. x j = p" with s a show "card ?S = 1" using ksimplex_replace_1[of p "n + 1" s a j] by (subst eq_commute) simp } { assume "card ?S \ 2" "\ (\j\{..n}. \x\s - {a}. x j = p)" with s a show "\j\{..n}. \x\s - {a}. x j = 0" using ksimplex_replace_2[of p "n + 1" s a] by (subst (asm) eq_commute) auto } qed subsubsection \Reduced labelling\ definition reduced :: "nat \ (nat \ nat) \ nat" where "reduced n x = (LEAST k. k = n \ x k \ 0)" lemma reduced_labelling: shows "reduced n x \ n" and "\i x (reduced n x) \ 0" proof - show "reduced n x \ n" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) auto show "\i x (reduced n x) \ 0" unfolding reduced_def by (rule LeastI2_wellorder[where a=n]) fastforce+ qed lemma reduced_labelling_unique: "r \ n \ \i r = n \ x r \ 0 \ reduced n x = r" by (metis linorder_less_linear linorder_not_le reduced_labelling) lemma reduced_labelling_zero: "j < n \ x j = 0 \ reduced n x \ j" using reduced_labelling[of n x] by auto lemma reduce_labelling_zero[simp]: "reduced 0 x = 0" by (rule reduced_labelling_unique) auto lemma reduced_labelling_nonzero: "j < n \ x j \ 0 \ reduced n x \ j" using reduced_labelling[of n x] by (elim allE[where x=j]) auto lemma reduced_labelling_Suc: "reduced (Suc n) x \ Suc n \ reduced (Suc n) x = reduced n x" using reduced_labelling[of "Suc n" x] by (intro reduced_labelling_unique[symmetric]) auto lemma complete_face_top: assumes "\x\f. \j\n. x j = 0 \ lab x j = 0" and "\x\f. \j\n. x j = p \ lab x j = 1" and eq: "(reduced (Suc n) \ lab) ` f = {..n}" shows "((\j\n. \x\f. x j = 0) \ (\j\n. \x\f. x j = p)) \ (\x\f. x n = p)" proof (safe del: disjCI) fix x j assume j: "j \ n" "\x\f. x j = 0" { fix x assume "x \ f" with assms j have "reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_zero) auto } moreover have "j \ (reduced (Suc n) \ lab) ` f" using j eq by auto ultimately show "x n = p" by force next fix x j assume j: "j \ n" "\x\f. x j = p" and x: "x \ f" have "j = n" proof (rule ccontr) assume "\ ?thesis" { fix x assume "x \ f" with assms j have "reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_nonzero) auto then have "reduced (Suc n) (lab x) \ n" using \j \ n\ \j \ n\ by simp } moreover have "n \ (reduced (Suc n) \ lab) ` f" using eq by auto ultimately show False by force qed moreover have "j \ (reduced (Suc n) \ lab) ` f" using j eq by auto ultimately show "x n = p" using j x by auto qed auto text \Hence we get just about the nice induction.\ lemma kuhn_induction: assumes "0 < p" and lab_0: "\x. \j\n. (\j. x j \ p) \ x j = 0 \ lab x j = 0" and lab_1: "\x. \j\n. (\j. x j \ p) \ x j = p \ lab x j = 1" and odd: "odd (card {s. ksimplex p n s \ (reduced n\lab) ` s = {..n}})" shows "odd (card {s. ksimplex p (Suc n) s \ (reduced (Suc n)\lab) ` s = {..Suc n}})" proof - let ?rl = "reduced (Suc n) \ lab" and ?ext = "\f v. \j\n. \x\f. x j = v" let ?ext = "\s. (\j\n. \x\s. x j = 0) \ (\j\n. \x\s. x j = p)" have "\s. ksimplex p (Suc n) s \ ?rl ` s \ {..Suc n}" by (simp add: reduced_labelling subset_eq) moreover have "{s. ksimplex p n s \ (reduced n \ lab) ` s = {..n}} = {f. \s a. ksimplex p (Suc n) s \ a \ s \ f = s - {a} \ ?rl ` f = {..n} \ ?ext f}" proof (intro set_eqI, safe del: disjCI equalityI disjE) fix s assume s: "ksimplex p n s" and rl: "(reduced n \ lab) ` s = {..n}" from s obtain u b where "kuhn_simplex p n u b s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p n u b s . have all_eq_p: "\x\s. x n = p" by (auto simp: out_eq_p) moreover { fix x assume "x \ s" with lab_1[rule_format, of n x] all_eq_p s_le_p[of x] have "?rl x \ n" by (auto intro!: reduced_labelling_nonzero) then have "?rl x = reduced n (lab x)" by (auto intro!: reduced_labelling_Suc) } then have "?rl ` s = {..n}" using rl by (simp cong: image_cong) moreover obtain t a where "ksimplex p (Suc n) t" "a \ t" "s = t - {a}" using s unfolding simplex_top_face[OF \0 < p\ all_eq_p] by auto ultimately show "\t a. ksimplex p (Suc n) t \ a \ t \ s = t - {a} \ ?rl ` s = {..n} \ ?ext s" by auto next fix x s a assume s: "ksimplex p (Suc n) s" and rl: "?rl ` (s - {a}) = {.. n}" and a: "a \ s" and "?ext (s - {a})" from s obtain u b where "kuhn_simplex p (Suc n) u b s" by (auto elim: ksimplex.cases) then interpret kuhn_simplex p "Suc n" u b s . have all_eq_p: "\x\s. x (Suc n) = p" by (auto simp: out_eq_p) { fix x assume "x \ s - {a}" then have "?rl x \ ?rl ` (s - {a})" by auto then have "?rl x \ n" unfolding rl by auto then have "?rl x = reduced n (lab x)" by (auto intro!: reduced_labelling_Suc) } then show rl': "(reduced n\lab) ` (s - {a}) = {..n}" unfolding rl[symmetric] by (intro image_cong) auto from \?ext (s - {a})\ have all_eq_p: "\x\s - {a}. x n = p" proof (elim disjE exE conjE) fix j assume "j \ n" "\x\s - {a}. x j = 0" with lab_0[rule_format, of j] all_eq_p s_le_p have "\x. x \ s - {a} \ reduced (Suc n) (lab x) \ j" by (intro reduced_labelling_zero) auto moreover have "j \ ?rl ` (s - {a})" using \j \ n\ unfolding rl by auto ultimately show ?thesis by force next fix j assume "j \ n" and eq_p: "\x\s - {a}. x j = p" show ?thesis proof cases assume "j = n" with eq_p show ?thesis by simp next assume "j \ n" { fix x assume x: "x \ s - {a}" have "reduced n (lab x) \ j" proof (rule reduced_labelling_nonzero) show "lab x j \ 0" using lab_1[rule_format, of j x] x s_le_p[of x] eq_p \j \ n\ by auto show "j < n" using \j \ n\ \j \ n\ by simp qed then have "reduced n (lab x) \ n" using \j \ n\ \j \ n\ by simp } moreover have "n \ (reduced n\lab) ` (s - {a})" unfolding rl' by auto ultimately show ?thesis by force qed qed show "ksimplex p n (s - {a})" unfolding simplex_top_face[OF \0 < p\ all_eq_p] using s a by auto qed ultimately show ?thesis using assms by (intro kuhn_simplex_lemma) auto qed text \And so we get the final combinatorial result.\ lemma ksimplex_0: "ksimplex p 0 s \ s = {(\x. p)}" proof assume "ksimplex p 0 s" then show "s = {(\x. p)}" by (blast dest: kuhn_simplex.ksimplex_0 elim: ksimplex.cases) next assume s: "s = {(\x. p)}" show "ksimplex p 0 s" proof (intro ksimplex, unfold_locales) show "(\_. p) \ {..<0::nat} \ {..x j. (\j. x j \ p) \ j < n \ x j = 0 \ lab x j = 0" and "\x j. (\j. x j \ p) \ j < n \ x j = p \ lab x j = 1" shows "odd (card {s. ksimplex p n s \ (reduced n\lab) ` s = {..n}})" (is "odd (card (?M n))") using assms proof (induct n) case 0 then show ?case by (simp add: ksimplex_0 cong: conj_cong) next case (Suc n) then have "odd (card (?M n))" by force with Suc show ?case using kuhn_induction[of p n] by (auto simp: comp_def) qed lemma kuhn_lemma: fixes n p :: nat assumes "0 < p" and "\x. (\i p) \ (\i label x i = 1)" and "\x. (\i p) \ (\i label x i = 0)" and "\x. (\i p) \ (\i label x i = 1)" obtains q where "\iir s. (\j r j \ r j \ q j + 1) \ (\j s j \ s j \ q j + 1) \ label r i \ label s i" proof - let ?rl = "reduced n \ label" let ?A = "{s. ksimplex p n s \ ?rl ` s = {..n}}" have "odd (card ?A)" using assms by (intro kuhn_combinatorial[of p n label]) auto then have "?A \ {}" by (rule odd_card_imp_not_empty) then obtain s b u where "kuhn_simplex p n b u s" and rl: "?rl ` s = {..n}" by (auto elim: ksimplex.cases) interpret kuhn_simplex p n b u s by fact show ?thesis proof (intro that[of b] allI impI) fix i assume "i < n" then show "b i < p" using base by auto next fix i assume "i < n" then have "i \ {.. n}" "Suc i \ {.. n}" by auto then obtain u v where u: "u \ s" "Suc i = ?rl u" and v: "v \ s" "i = ?rl v" unfolding rl[symmetric] by blast have "label u i \ label v i" using reduced_labelling [of n "label u"] reduced_labelling [of n "label v"] u(2)[symmetric] v(2)[symmetric] \i < n\ by auto moreover have "b j \ u j" "u j \ b j + 1" "b j \ v j" "v j \ b j + 1" if "j < n" for j using that base_le[OF \u\s\] le_Suc_base[OF \u\s\] base_le[OF \v\s\] le_Suc_base[OF \v\s\] by auto ultimately show "\r s. (\j r j \ r j \ b j + 1) \ (\j s j \ s j \ b j + 1) \ label r i \ label s i" by blast qed qed subsubsection \Main result for the unit cube\ lemma kuhn_labelling_lemma': assumes "(\x::nat\real. P x \ P (f x))" and "\x. P x \ (\i::nat. Q i \ 0 \ x i \ x i \ 1)" shows "\l. (\x i. l x i \ (1::nat)) \ (\x i. P x \ Q i \ x i = 0 \ l x i = 0) \ (\x i. P x \ Q i \ x i = 1 \ l x i = 1) \ (\x i. P x \ Q i \ l x i = 0 \ x i \ f x i) \ (\x i. P x \ Q i \ l x i = 1 \ f x i \ x i)" unfolding all_conj_distrib [symmetric] apply (subst choice_iff[symmetric])+ - by (metis assms bot_nat_0.extremum nle_le zero_neq_one) + by (metis assms choice_iff bot_nat_0.extremum nle_le zero_neq_one) subsection \Brouwer's fixed point theorem\ text \We start proving Brouwer's fixed point theorem for the unit cube = \cbox 0 One\.\ lemma brouwer_cube: fixes f :: "'a::euclidean_space \ 'a" assumes "continuous_on (cbox 0 One) f" and "f ` cbox 0 One \ cbox 0 One" shows "\x\cbox 0 One. f x = x" proof (rule ccontr) define n where "n = DIM('a)" have n: "1 \ n" "0 < n" "n \ 0" unfolding n_def by (auto simp: Suc_le_eq) assume "\ ?thesis" then have *: "\ (\x\cbox 0 One. f x - x = 0)" by auto obtain d where d: "d > 0" "\x. x \ cbox 0 One \ d \ norm (f x - x)" - apply (rule brouwer_compactness_lemma[OF compact_cbox _ *]) - apply (rule continuous_intros assms)+ - apply blast - done + using brouwer_compactness_lemma[OF compact_cbox _ *] assms + by (metis (no_types, lifting) continuous_on_cong continuous_on_diff continuous_on_id) have *: "\x. x \ cbox 0 One \ f x \ cbox 0 One" "\x. x \ (cbox 0 One::'a set) \ (\i\Basis. True \ 0 \ x \ i \ x \ i \ 1)" using assms(2)[unfolded image_subset_iff Ball_def] unfolding cbox_def by auto obtain label :: "'a \ 'a \ nat" where label [rule_format]: "\x. \i\Basis. label x i \ 1" "\x. \i\Basis. x \ cbox 0 One \ x \ i = 0 \ label x i = 0" "\x. \i\Basis. x \ cbox 0 One \ x \ i = 1 \ label x i = 1" "\x. \i\Basis. x \ cbox 0 One \ label x i = 0 \ x \ i \ f x \ i" "\x. \i\Basis. x \ cbox 0 One \ label x i = 1 \ f x \ i \ x \ i" using kuhn_labelling_lemma[OF *] by auto note label = this [rule_format] have lem1: "\x\cbox 0 One. \y\cbox 0 One. \i\Basis. label x i \ label y i \ \f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" proof safe fix x y :: 'a assume x: "x \ cbox 0 One" and y: "y \ cbox 0 One" fix i assume i: "label x i \ label y i" "i \ Basis" have *: "\x y fx fy :: real. x \ fx \ fy \ y \ fx \ x \ y \ fy \ \fx - x\ \ \fy - fx\ + \y - x\" by auto have "\(f x - x) \ i\ \ \(f y - f x)\i\ + \(y - x)\i\" proof (cases "label x i = 0") case True then have fxy: "\ f y \ i \ y \ i \ f x \ i \ x \ i" by (metis True i label(1) label(5) le_antisym less_one not_le_imp_less y) show ?thesis unfolding inner_simps by (rule *) (auto simp: True i label x y fxy) next case False then show ?thesis using label [OF \i \ Basis\] i(1) x y - apply (auto simp: inner_diff_left le_Suc_eq) - by (metis "*") + by (smt (verit, ccfv_threshold) inner_diff_left less_one order_le_less) qed also have "\ \ norm (f y - f x) + norm (y - x)" by (simp add: add_mono i(2) norm_bound_Basis_le) finally show "\f x \ i - x \ i\ \ norm (f y - f x) + norm (y - x)" unfolding inner_simps . qed have "\e>0. \x\cbox 0 One. \y\cbox 0 One. \z\cbox 0 One. \i\Basis. norm (x - z) < e \ norm (y - z) < e \ label x i \ label y i \ \(f(z) - z)\i\ < d / (real n)" proof - have d': "d / real n / 8 > 0" using d(1) by (simp add: n_def) have *: "uniformly_continuous_on (cbox 0 One) f" by (rule compact_uniformly_continuous[OF assms(1) compact_cbox]) obtain e where e: "e > 0" "\x x'. x \ cbox 0 One \ x' \ cbox 0 One \ norm (x' - x) < e \ norm (f x' - f x) < d / real n / 8" using *[unfolded uniformly_continuous_on_def,rule_format,OF d'] unfolding dist_norm by blast show ?thesis proof (intro exI conjI ballI impI) show "0 < min (e / 2) (d / real n / 8)" using d' e by auto fix x y z i assume as: "x \ cbox 0 One" "y \ cbox 0 One" "z \ cbox 0 One" "norm (x - z) < min (e / 2) (d / real n / 8)" "norm (y - z) < min (e / 2) (d / real n / 8)" "label x i \ label y i" assume i: "i \ Basis" have *: "\z fz x fx n1 n2 n3 n4 d4 d :: real. \fx - x\ \ n1 + n2 \ \fx - fz\ \ n3 \ \x - z\ \ n4 \ n1 < d4 \ n2 < 2 * d4 \ n3 < d4 \ n4 < d4 \ (8 * d4 = d) \ \fz - z\ < d" by auto show "\(f z - z) \ i\ < d / real n" unfolding inner_simps proof (rule *) show "\f x \ i - x \ i\ \ norm (f y -f x) + norm (y - x)" using as(1) as(2) as(6) i lem1 by blast show "norm (f x - f z) < d / real n / 8" using d' e as by auto show "\f x \ i - f z \ i\ \ norm (f x - f z)" "\x \ i - z \ i\ \ norm (x - z)" unfolding inner_diff_left[symmetric] by (rule Basis_le_norm[OF i])+ have tria: "norm (y - x) \ norm (y - z) + norm (x - z)" using dist_triangle[of y x z, unfolded dist_norm] unfolding norm_minus_commute by auto also have "\ < e / 2 + e / 2" using as(4) as(5) by auto finally show "norm (f y - f x) < d / real n / 8" using as(1) as(2) e(2) by auto have "norm (y - z) + norm (x - z) < d / real n / 8 + d / real n / 8" using as(4) as(5) by auto with tria show "norm (y - x) < 2 * (d / real n / 8)" by auto qed (use as in auto) qed qed then obtain e where e: "e > 0" "\x y z i. x \ cbox 0 One \ y \ cbox 0 One \ z \ cbox 0 One \ i \ Basis \ norm (x - z) < e \ norm (y - z) < e \ label x i \ label y i \ \(f z - z) \ i\ < d / real n" by blast obtain p :: nat where p: "1 + real n / e \ real p" using real_arch_simple .. have "1 + real n / e > 0" using e(1) n by (simp add: add_pos_pos) then have "p > 0" using p by auto obtain b :: "nat \ 'a" where b: "bij_betw b {..< n} Basis" by atomize_elim (auto simp: n_def intro!: finite_same_card_bij) define b' where "b' = inv_into {..< n} b" then have b': "bij_betw b' Basis {..< n}" using bij_betw_inv_into[OF b] by auto then have b'_Basis: "\i. i \ Basis \ b' i \ {..< n}" unfolding bij_betw_def by (auto simp: set_eq_iff) have bb'[simp]:"\i. i \ Basis \ b (b' i) = i" unfolding b'_def using b by (auto simp: f_inv_into_f bij_betw_def) have b'b[simp]:"\i. i < n \ b' (b i) = i" unfolding b'_def using b by (auto simp: inv_into_f_eq bij_betw_def) have *: "\x :: nat. x = 0 \ x = 1 \ x \ 1" by auto have b'': "\j. j < n \ b j \ Basis" using b unfolding bij_betw_def by auto have q1: "0 < p" "\x. (\i p) \ (\ii\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0 \ (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" unfolding * using \p > 0\ \n > 0\ using label(1)[OF b''] by auto { fix x :: "nat \ nat" and i assume "\i p" "i < n" "x i = p \ x i = 0" then have "(\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ (cbox 0 One::'a set)" using b'_Basis by (auto simp: cbox_def bij_betw_def zero_le_divide_iff divide_le_eq_1) } note cube = this have q2: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 0)" unfolding o_def using cube \p > 0\ by (intro allI impI label(2)) (auto simp: b'') have q3: "\x. (\i p) \ (\i (label (\i\Basis. (real (x (b' i)) / real p) *\<^sub>R i) \ b) i = 1)" using cube \p > 0\ unfolding o_def by (intro allI impI label(3)) (auto simp: b'') obtain q where q: "\iir s. (\j r j \ r j \ q j + 1) \ (\j s j \ s j \ q j + 1) \ (label (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i) \ b) i \ (label (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i) \ b) i" by (rule kuhn_lemma[OF q1 q2 q3]) define z :: 'a where "z = (\i\Basis. (real (q (b' i)) / real p) *\<^sub>R i)" have "\i\Basis. d / real n \ \(f z - z)\i\" proof (rule ccontr) have "\i\Basis. q (b' i) \ {0..p}" using q(1) b' by (auto intro: less_imp_le simp: bij_betw_def) then have "z \ cbox 0 One" unfolding z_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) then have d_fz_z: "d \ norm (f z - z)" by (rule d) assume "\ ?thesis" then have as: "\i\Basis. \f z \ i - z \ i\ < d / real n" using \n > 0\ by (auto simp: not_le inner_diff) have "norm (f z - z) \ (\i\Basis. \f z \ i - z \ i\)" unfolding inner_diff_left[symmetric] by (rule norm_le_l1) also have "\ < (\(i::'a) \ Basis. d / real n)" by (meson as finite_Basis nonempty_Basis sum_strict_mono) also have "\ = d" using DIM_positive[where 'a='a] by (auto simp: n_def) finally show False using d_fz_z by auto qed then obtain i where i: "i \ Basis" "d / real n \ \(f z - z) \ i\" .. have *: "b' i < n" using i and b'[unfolded bij_betw_def] by auto obtain r s where rs: "\j. j < n \ q j \ r j \ r j \ q j + 1" "\j. j < n \ q j \ s j \ s j \ q j + 1" "(label (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i) \ b) (b' i) \ (label (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i) \ b) (b' i)" using q(2)[rule_format,OF *] by blast have b'_im: "\i. i \ Basis \ b' i < n" using b' unfolding bij_betw_def by auto define r' ::'a where "r' = (\i\Basis. (real (r (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ r (b' i) \ p" using b'_im q(1) rs(1) by fastforce then have "r' \ cbox 0 One" unfolding r'_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) define s' :: 'a where "s' = (\i\Basis. (real (s (b' i)) / real p) *\<^sub>R i)" have "\i. i \ Basis \ s (b' i) \ p" using b'_im q(1) rs(2) by fastforce then have "s' \ cbox 0 One" unfolding s'_def cbox_def using b'_Basis by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1) have "z \ cbox 0 One" unfolding z_def cbox_def using b'_Basis q(1)[rule_format,OF b'_im] \p > 0\ by (auto simp: bij_betw_def zero_le_divide_iff divide_le_eq_1 less_imp_le) { have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" by (rule sum_mono) (use rs(1)[OF b'_im] in force) also have "\ < e * real p" using p \e > 0\ \p > 0\ by (auto simp: field_simps n_def) finally have "(\i\Basis. \real (r (b' i)) - real (q (b' i))\) < e * real p" . } moreover { have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) \ (\(i::'a)\Basis. 1)" by (rule sum_mono) (use rs(2)[OF b'_im] in force) also have "\ < e * real p" using p \e > 0\ \p > 0\ by (auto simp: field_simps n_def) finally have "(\i\Basis. \real (s (b' i)) - real (q (b' i))\) < e * real p" . } ultimately have "norm (r' - z) < e" and "norm (s' - z) < e" unfolding r'_def s'_def z_def using \p > 0\ apply (rule_tac[!] le_less_trans[OF norm_le_l1]) apply (auto simp: field_simps sum_divide_distrib[symmetric] inner_diff_left) done then have "\(f z - z) \ i\ < d / real n" using rs(3) i unfolding r'_def[symmetric] s'_def[symmetric] o_def bb' by (intro e(2)[OF \r'\cbox 0 One\ \s'\cbox 0 One\ \z\cbox 0 One\]) auto then show False using i by auto qed text \Next step is to prove it for nonempty interiors.\ lemma brouwer_weak: fixes f :: "'a::euclidean_space \ 'a" assumes "compact S" and "convex S" and "interior S \ {}" and "continuous_on S f" and "f \ S \ S" obtains x where "x \ S" and "f x = x" proof - let ?U = "cbox 0 One :: 'a set" have "\Basis /\<^sub>R 2 \ interior ?U" proof (rule interiorI) let ?I = "(\i\Basis. {x::'a. 0 < x \ i} \ {x. x \ i < 1})" show "open ?I" by (intro open_INT finite_Basis ballI open_Int, auto intro: open_Collect_less simp: continuous_on_inner) show "\Basis /\<^sub>R 2 \ ?I" by simp show "?I \ cbox 0 One" unfolding cbox_def by force qed then have *: "interior ?U \ {}" by fast have *: "?U homeomorphic S" using homeomorphic_convex_compact[OF convex_box(1) compact_cbox * assms(2,1,3)] . have "\f. continuous_on ?U f \ f \ ?U \ ?U \ (\x\?U. f x = x)" using brouwer_cube by auto then show ?thesis unfolding homeomorphic_fixpoint_property[OF *] using assms by (auto intro: that) qed text \Then the particular case for closed balls.\ lemma brouwer_ball: fixes f :: "'a::euclidean_space \ 'a" assumes "e > 0" and "continuous_on (cball a e) f" and "f \ cball a e \ cball a e" obtains x where "x \ cball a e" and "f x = x" using brouwer_weak[OF compact_cball convex_cball, of a e f] unfolding interior_cball ball_eq_empty using assms by auto text \And finally we prove Brouwer's fixed point theorem in its general version.\ theorem brouwer: fixes f :: "'a::euclidean_space \ 'a" assumes S: "compact S" "convex S" "S \ {}" and contf: "continuous_on S f" and fim: "f \ S \ S" obtains x where "x \ S" and "f x = x" proof - have "\e>0. S \ cball 0 e" using compact_imp_bounded[OF \compact S\] unfolding bounded_pos by auto then obtain e where e: "e > 0" "S \ cball 0 e" by blast have "\x\ cball 0 e. (f \ closest_point S) x = x" proof (rule_tac brouwer_ball[OF e(1)]) show "continuous_on (cball 0 e) (f \ closest_point S)" by (meson assms closest_point_in_set compact_eq_bounded_closed contf continuous_on_closest_point continuous_on_compose continuous_on_subset image_subsetI) show "f \ closest_point S \ cball 0 e \ cball 0 e" by (smt (verit) Pi_iff assms(1) assms(3) closest_point_in_set comp_apply compact_eq_bounded_closed e(2) fim subset_eq) qed (use assms in auto) then obtain x where x: "x \ cball 0 e" "(f \ closest_point S) x = x" .. with S have "x \ S" by (metis PiE closest_point_in_set comp_apply compact_imp_closed fim) then have *: "closest_point S x = x" by (rule closest_point_self) show thesis proof show "closest_point S x \ S" by (simp add: "*" \x \ S\) show "f (closest_point S x) = closest_point S x" using "*" x by auto qed qed subsection \Applications\ text \So we get the no-retraction theorem.\ corollary no_retraction_cball: fixes a :: "'a::euclidean_space" assumes "e > 0" shows "\ (frontier (cball a e) retract_of (cball a e))" proof assume *: "frontier (cball a e) retract_of (cball a e)" have **: "\xa. a - (2 *\<^sub>R a - xa) = - (a - xa)" using scaleR_left_distrib[of 1 1 a] by auto obtain x where x: "x \ {x. norm (a - x) = e}" "2 *\<^sub>R a - x = x" proof (rule retract_fixpoint_property[OF *, of "\x. scaleR 2 a - x"]) show "continuous_on (frontier (cball a e)) ((-) (2 *\<^sub>R a))" by (intro continuous_intros) show "(-) (2 *\<^sub>R a) \ frontier (cball a e) \ frontier (cball a e)" by clarsimp (metis "**" dist_norm norm_minus_cancel) qed (auto simp: dist_norm intro: brouwer_ball[OF assms]) then have "scaleR 2 a = scaleR 1 x + scaleR 1 x" by (auto simp: algebra_simps) then have "a = x" unfolding scaleR_left_distrib[symmetric] by auto then show False using x assms by auto qed corollary contractible_sphere: fixes a :: "'a::euclidean_space" shows "contractible(sphere a r) \ r \ 0" proof (cases "0 < r") case True then show ?thesis unfolding contractible_def nullhomotopic_from_sphere_extension using no_retraction_cball [OF True, of a] by (auto simp: retract_of_def retraction_def) next case False then show ?thesis unfolding contractible_def nullhomotopic_from_sphere_extension using less_eq_real_def by auto qed corollary connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by auto next case equal then show ?thesis by auto next case greater show ?thesis proof assume L: ?lhs have "False" if 1: "DIM('a) = 1" proof - obtain x y where xy: "sphere a r = {x,y}" "x \ y" using sphere_1D_doubleton [OF 1 greater] by (metis dist_self greater insertI1 less_add_same_cancel1 mem_sphere mult_2 not_le zero_le_dist) then have "finite (sphere a r)" by auto with L \r > 0\ xy show "False" using connected_finite_iff_sing by auto qed with greater show ?rhs by (metis DIM_ge_Suc0 One_nat_def Suc_1 le_antisym not_less_eq_eq) next assume ?rhs then show ?lhs using connected_sphere greater by auto qed qed corollary path_connected_sphere_eq: fixes a :: "'a :: euclidean_space" shows "path_connected(sphere a r) \ 2 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using connected_sphere_eq path_connected_imp_connected by blast next assume R: ?rhs then show ?lhs by (auto simp: contractible_imp_path_connected contractible_sphere path_connected_sphere) qed proposition frontier_subset_retraction: fixes S :: "'a::euclidean_space set" assumes "bounded S" and fros: "frontier S \ T" and contf: "continuous_on (closure S) f" and fim: "f \ S \ T" and fid: "\x. x \ T \ f x = x" shows "S \ T" proof (rule ccontr) assume "\ S \ T" then obtain a where "a \ S" "a \ T" by blast define g where "g \ \z. if z \ closure S then f z else z" have "continuous_on (closure S \ closure(-S)) g" - unfolding g_def - apply (rule continuous_on_cases) - using fros fid frontier_closures by (auto simp: contf) + unfolding g_def using fros fid frontier_closures + by (intro continuous_on_cases) (auto simp: contf) moreover have "closure S \ closure(- S) = UNIV" using closure_Un by fastforce ultimately have contg: "continuous_on UNIV g" by metis obtain B where "0 < B" and B: "closure S \ ball a B" using \bounded S\ bounded_subset_ballD by blast have notga: "g x \ a" for x unfolding g_def using fros fim \a \ T\ by (metis PiE Un_iff \a \ S\ closure_Un_frontier fid subsetD) define h where "h \ (\y. a + (B / norm(y - a)) *\<^sub>R (y - a)) \ g" have "\ (frontier (cball a B) retract_of (cball a B))" by (metis no_retraction_cball \0 < B\) then have "\k. \ retraction (cball a B) (frontier (cball a B)) k" by (simp add: retract_of_def) moreover have "retraction (cball a B) (frontier (cball a B)) h" unfolding retraction_def proof (intro conjI ballI) show "frontier (cball a B) \ cball a B" by force show "continuous_on (cball a B) h" unfolding h_def by (intro continuous_intros) (use contg continuous_on_subset notga in auto) show "h \ cball a B \ frontier (cball a B)" using \0 < B\ by (auto simp: h_def notga dist_norm) show "\x. x \ frontier (cball a B) \ h x = x" using notga \0 < B\ apply (simp add: g_def h_def field_simps) by (metis B dist_commute dist_norm mem_ball order_less_irrefl subset_eq) qed ultimately show False by simp qed subsubsection \Punctured affine hulls, etc\ lemma rel_frontier_deformation_retract_of_punctured_convex: fixes S :: "'a::euclidean_space set" assumes "convex S" "convex T" "bounded S" and arelS: "a \ rel_interior S" and relS: "rel_frontier S \ T" and affS: "T \ affine hull S" obtains r where "homotopic_with_canon (\x. True) (T - {a}) (T - {a}) id r" "retraction (T - {a}) (rel_frontier S) r" proof - have "\d. 0 < d \ (a + d *\<^sub>R l) \ rel_frontier S \ (\e. 0 \ e \ e < d \ (a + e *\<^sub>R l) \ rel_interior S)" - if "(a + l) \ affine hull S" "l \ 0" for l - apply (rule ray_to_rel_frontier [OF \bounded S\ arelS]) - apply (rule that)+ - by metis + if "(a + l) \ affine hull S" "l \ 0" for l + using ray_to_rel_frontier [OF \bounded S\ arelS] that by metis then obtain dd where dd1: "\l. \(a + l) \ affine hull S; l \ 0\ \ 0 < dd l \ (a + dd l *\<^sub>R l) \ rel_frontier S" and dd2: "\l e. \(a + l) \ affine hull S; e < dd l; 0 \ e; l \ 0\ \ (a + e *\<^sub>R l) \ rel_interior S" by metis+ have aaffS: "a \ affine hull S" by (meson arelS subsetD hull_inc rel_interior_subset) have "((\z. z - a) ` (affine hull S - {a})) = ((\z. z - a) ` (affine hull S)) - {0}" by auto moreover have "continuous_on (((\z. z - a) ` (affine hull S)) - {0}) (\x. dd x *\<^sub>R x)" proof (rule continuous_on_compact_surface_projection) show "compact (rel_frontier ((\z. z - a) ` S))" by (simp add: \bounded S\ bounded_translation_minus compact_rel_frontier_bounded) have releq: "rel_frontier ((\z. z - a) ` S) = (\z. z - a) ` rel_frontier S" using rel_frontier_translation [of "-a"] add.commute by simp also have "\ \ (\z. z - a) ` (affine hull S) - {0}" using rel_frontier_affine_hull arelS rel_frontier_def by fastforce finally show "rel_frontier ((\z. z - a) ` S) \ (\z. z - a) ` (affine hull S) - {0}" . show "cone ((\z. z - a) ` (affine hull S))" by (rule subspace_imp_cone) (use aaffS in \simp add: subspace_affine image_comp o_def affine_translation_aux [of a]\) show "(0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)) \ (dd x = k)" if x: "x \ (\z. z - a) ` (affine hull S) - {0}" for k x proof show "dd x = k \ 0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)" using dd1 [of x] that image_iff by (fastforce simp add: releq) next assume k: "0 < k \ k *\<^sub>R x \ rel_frontier ((\z. z - a) ` S)" have False if "dd x < k" proof - have "k \ 0" "a + k *\<^sub>R x \ closure S" using k closure_translation [of "-a"] by (auto simp: rel_frontier_def cong: image_cong_simp) then have segsub: "open_segment a (a + k *\<^sub>R x) \ rel_interior S" by (metis rel_interior_closure_convex_segment [OF \convex S\ arelS]) have "x \ 0" and xaffS: "a + x \ affine hull S" using x by auto then have "0 < dd x" and inS: "a + dd x *\<^sub>R x \ rel_frontier S" using dd1 by auto moreover have "a + dd x *\<^sub>R x \ open_segment a (a + k *\<^sub>R x)" - using k \x \ 0\ \0 < dd x\ - apply (simp add: in_segment) - apply (rule_tac x = "dd x / k" in exI) - apply (simp add: that vector_add_divide_simps algebra_simps) - done + unfolding in_segment + proof (intro conjI exI) + show "a + dd x *\<^sub>R x = (1 - dd x / k) *\<^sub>R a + (dd x / k) *\<^sub>R (a + k *\<^sub>R x)" + using k by (simp add: that algebra_simps) + qed (use \x \ 0\ \0 < dd x\ that in auto) ultimately show ?thesis using segsub by (auto simp: rel_frontier_def) qed moreover have False if "k < dd x" using x k that rel_frontier_def by (fastforce simp: algebra_simps releq dest!: dd2) ultimately show "dd x = k" by fastforce qed qed ultimately have *: "continuous_on ((\z. z - a) ` (affine hull S - {a})) (\x. dd x *\<^sub>R x)" by auto have "continuous_on (affine hull S - {a}) ((\x. a + dd x *\<^sub>R x) \ (\z. z - a))" by (intro * continuous_intros continuous_on_compose) with affS have contdd: "continuous_on (T - {a}) ((\x. a + dd x *\<^sub>R x) \ (\z. z - a))" by (blast intro: continuous_on_subset) show ?thesis proof - show "homotopic_with_canon (\x. True) (T - {a}) (T - {a}) id (\x. a + dd (x - a) *\<^sub>R (x - a))" + show "homotopic_with_canon (\x. True) (T - {a}) (T - {a}) id (\x. a + dd (x-a) *\<^sub>R (x-a))" proof (rule homotopic_with_linear) show "continuous_on (T - {a}) id" by (intro continuous_intros continuous_on_compose) - show "continuous_on (T - {a}) (\x. a + dd (x - a) *\<^sub>R (x - a))" + show "continuous_on (T - {a}) (\x. a + dd (x-a) *\<^sub>R (x-a))" using contdd by (simp add: o_def) - show "closed_segment (id x) (a + dd (x - a) *\<^sub>R (x - a)) \ T - {a}" + show "closed_segment (id x) (a + dd (x-a) *\<^sub>R (x-a)) \ T - {a}" if "x \ T - {a}" for x proof (clarsimp simp: in_segment, intro conjI) fix u::real assume u: "0 \ u" "u \ 1" - have "a + dd (x - a) *\<^sub>R (x - a) \ T" + have "a + dd (x-a) *\<^sub>R (x-a) \ T" by (metis DiffD1 DiffD2 add.commute add.right_neutral affS dd1 diff_add_cancel relS singletonI subsetCE that) - then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \ T" + then show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x-a) *\<^sub>R (x-a)) \ T" using convexD [OF \convex T\] that u by simp - have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x - a)) = a \ - (1 - u + u * d) *\<^sub>R (x - a) = 0" for d + have iff: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + d *\<^sub>R (x-a)) = a \ + (1 - u + u * d) *\<^sub>R (x-a) = 0" for d by (auto simp: algebra_simps) have "x \ T" "x \ a" using that by auto - then have axa: "a + (x - a) \ affine hull S" + then have axa: "a + (x-a) \ affine hull S" by (metis (no_types) add.commute affS diff_add_cancel rev_subsetD) - then have "\ dd (x - a) \ 0 \ a + dd (x - a) *\<^sub>R (x - a) \ rel_frontier S" + then have "\ dd (x-a) \ 0 \ a + dd (x-a) *\<^sub>R (x-a) \ rel_frontier S" using \x \ a\ dd1 by fastforce - with \x \ a\ show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x - a) *\<^sub>R (x - a)) \ a" + with \x \ a\ show "(1 - u) *\<^sub>R x + u *\<^sub>R (a + dd (x-a) *\<^sub>R (x-a)) \ a" using less_eq_real_def mult_le_0_iff not_less u by (fastforce simp: iff) qed qed - show "retraction (T - {a}) (rel_frontier S) (\x. a + dd (x - a) *\<^sub>R (x - a))" + show "retraction (T - {a}) (rel_frontier S) (\x. a + dd (x-a) *\<^sub>R (x-a))" proof (simp add: retraction_def, intro conjI ballI) show "rel_frontier S \ T - {a}" using arelS relS rel_frontier_def by fastforce - show "continuous_on (T - {a}) (\x. a + dd (x - a) *\<^sub>R (x - a))" + show "continuous_on (T - {a}) (\x. a + dd (x-a) *\<^sub>R (x-a))" using contdd by (simp add: o_def) - show "(\x. a + dd (x - a) *\<^sub>R (x - a)) \ (T - {a}) \ rel_frontier S" - apply (auto simp: rel_frontier_def) - apply (metis Diff_subset add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def subset_iff) - by (metis DiffE add.commute affS dd1 diff_add_cancel eq_iff_diff_eq_0 rel_frontier_def rev_subsetD) - show "a + dd (x - a) *\<^sub>R (x - a) = x" if x: "x \ rel_frontier S" for x + show "(\x. a + dd (x-a) *\<^sub>R (x-a)) \ (T - {a}) \ rel_frontier S" + unfolding Pi_iff using affS dd1 subset_eq by force + show "a + dd (x-a) *\<^sub>R (x-a) = x" if x: "x \ rel_frontier S" for x proof - have "x \ a" using that arelS by (auto simp: rel_frontier_def) - have False if "dd (x - a) < 1" + have False if "dd (x-a) < 1" proof - have "x \ closure S" using x by (auto simp: rel_frontier_def) then have segsub: "open_segment a x \ rel_interior S" by (metis rel_interior_closure_convex_segment [OF \convex S\ arelS]) have xaffS: "x \ affine hull S" using affS relS x by auto - then have "0 < dd (x - a)" and inS: "a + dd (x - a) *\<^sub>R (x - a) \ rel_frontier S" + then have "0 < dd (x-a)" and inS: "a + dd (x-a) *\<^sub>R (x-a) \ rel_frontier S" using dd1 by (auto simp: \x \ a\) - moreover have "a + dd (x - a) *\<^sub>R (x - a) \ open_segment a x" - using \x \ a\ \0 < dd (x - a)\ - apply (simp add: in_segment) - apply (rule_tac x = "dd (x - a)" in exI) - apply (simp add: algebra_simps that) - done + moreover have "a + dd (x-a) *\<^sub>R (x-a) \ open_segment a x" + unfolding in_segment + proof (intro exI conjI) + show "a + dd (x-a) *\<^sub>R (x-a) = (1 - dd (x-a)) *\<^sub>R a + (dd (x-a)) *\<^sub>R x" + by (simp add: algebra_simps) + qed (use \x \ a\ \0 < dd (x-a)\ that in auto) ultimately show ?thesis using segsub by (auto simp: rel_frontier_def) qed - moreover have False if "1 < dd (x - a)" + moreover have False if "1 < dd (x-a)" using x that dd2 [of "x - a" 1] \x \ a\ closure_affine_hull by (auto simp: rel_frontier_def) - ultimately have "dd (x - a) = 1" \ \similar to another proof above\ + ultimately have "dd (x-a) = 1" \ \similar to another proof above\ by fastforce with that show ?thesis by (simp add: rel_frontier_def) qed qed qed qed corollary rel_frontier_retract_of_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "bounded S" "convex S" "a \ rel_interior S" shows "rel_frontier S retract_of (affine hull S - {a})" by (meson assms convex_affine_hull dual_order.refl rel_frontier_affine_hull rel_frontier_deformation_retract_of_punctured_convex retract_of_def) corollary rel_boundary_retract_of_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "compact S" "convex S" "a \ rel_interior S" shows "(S - rel_interior S) retract_of (affine hull S - {a})" by (metis assms closure_closed compact_eq_bounded_closed rel_frontier_def rel_frontier_retract_of_punctured_affine_hull) lemma homotopy_eqv_rel_frontier_punctured_convex: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ rel_interior S" "convex T" "rel_frontier S \ T" "T \ affine hull S" shows "(rel_frontier S) homotopy_eqv (T - {a})" by (meson assms deformation_retract_imp_homotopy_eqv homotopy_equivalent_space_sym rel_frontier_deformation_retract_of_punctured_convex[of S T]) lemma homotopy_eqv_rel_frontier_punctured_affine_hull: fixes S :: "'a::euclidean_space set" assumes "convex S" "bounded S" "a \ rel_interior S" shows "(rel_frontier S) homotopy_eqv (affine hull S - {a})" by (simp add: assms homotopy_eqv_rel_frontier_punctured_convex rel_frontier_affine_hull) lemma path_connected_sphere_gen: assumes "convex S" "bounded S" "aff_dim S \ 1" shows "path_connected(rel_frontier S)" proof - have "convex (closure S)" using assms by auto then show ?thesis by (metis Diff_empty aff_dim_affine_hull assms convex_affine_hull convex_imp_path_connected equals0I path_connected_punctured_convex rel_frontier_def rel_frontier_retract_of_punctured_affine_hull retract_of_path_connected) qed lemma connected_sphere_gen: assumes "convex S" "bounded S" "aff_dim S \ 1" shows "connected(rel_frontier S)" by (simp add: assms path_connected_imp_connected path_connected_sphere_gen) subsubsection\Borsuk-style characterization of separation\ lemma continuous_on_Borsuk_map: - "a \ S \ continuous_on S (\x. inverse(norm (x - a)) *\<^sub>R (x - a))" + "a \ S \ continuous_on S (\x. inverse(norm (x-a)) *\<^sub>R (x-a))" by (rule continuous_intros | force)+ lemma Borsuk_map_into_sphere: - "(\x. inverse(norm (x - a)) *\<^sub>R (x - a)) \ S \ sphere 0 1 \ (a \ S)" + "(\x. inverse(norm (x-a)) *\<^sub>R (x-a)) \ S \ sphere 0 1 \ (a \ S)" proof - - have "\x. \a \ S; x \ S\ \ inverse (norm (x - a)) * norm (x - a) = 1" + have "\x. \a \ S; x \ S\ \ inverse (norm (x-a)) * norm (x-a) = 1" by (metis left_inverse norm_eq_zero right_minus_eq) then show ?thesis by force qed lemma Borsuk_maps_homotopic_in_path_component: assumes "path_component (- S) a b" shows "homotopic_with_canon (\x. True) S (sphere 0 1) - (\x. inverse(norm(x - a)) *\<^sub>R (x - a)) + (\x. inverse(norm(x-a)) *\<^sub>R (x-a)) (\x. inverse(norm(x - b)) *\<^sub>R (x - b))" proof - - obtain g where "path g" "path_image g \ -S" "pathstart g = a" "pathfinish g = b" + obtain g where g: "path g" "path_image g \ -S" "pathstart g = a" "pathfinish g = b" using assms by (auto simp: path_component_def) - then show ?thesis - apply (simp add: path_def path_image_def pathstart_def pathfinish_def homotopic_with_def) - apply (rule_tac x = "\z. inverse(norm(snd z - (g \ fst)z)) *\<^sub>R (snd z - (g \ fst)z)" in exI) - apply (rule conjI continuous_intros | erule continuous_on_subset | fastforce simp: divide_simps sphere_def)+ - done + define h where "h \ \z. (snd z - (g \ fst) z) /\<^sub>R norm (snd z - (g \ fst) z)" + have "continuous_on ({0..1} \ S) h" + unfolding h_def using g by (intro continuous_intros) (auto simp: path_defs) + moreover + have "h ` ({0..1} \ S) \ sphere 0 1" + unfolding h_def using g by (auto simp: divide_simps path_defs) + ultimately show ?thesis + using g by (auto simp: h_def path_defs homotopic_with_def) qed lemma non_extensible_Borsuk_map: fixes a :: "'a :: euclidean_space" assumes "compact S" and cin: "C \ components(- S)" and boc: "bounded C" and "a \ C" shows "\ (\g. continuous_on (S \ C) g \ g \ (S \ C) \ sphere 0 1 \ - (\x \ S. g x = inverse(norm(x - a)) *\<^sub>R (x - a)))" + (\x \ S. g x = inverse(norm(x-a)) *\<^sub>R (x-a)))" proof - have "closed S" using assms by (simp add: compact_imp_closed) have "C \ -S" using assms by (simp add: in_components_subset) with \a \ C\ have "a \ S" by blast then have ceq: "C = connected_component_set (- S) a" by (metis \a \ C\ cin components_iff connected_component_eq) then have "bounded (S \ connected_component_set (- S) a)" using \compact S\ boc compact_imp_bounded by auto with bounded_subset_ballD obtain r where "0 < r" and r: "(S \ connected_component_set (- S) a) \ ball a r" by blast { fix g assume "continuous_on (S \ C) g" "g \ (S \ C) \ sphere 0 1" - and [simp]: "\x. x \ S \ g x = (x - a) /\<^sub>R norm (x - a)" + and [simp]: "\x. x \ S \ g x = (x-a) /\<^sub>R norm (x-a)" then have norm_g1[simp]: "\x. x \ S \ C \ norm (g x) = 1" by force have cb_eq: "cball a r = (S \ connected_component_set (- S) a) \ (cball a r - connected_component_set (- S) a)" using ball_subset_cball [of a r] r by auto have cont1: "continuous_on (S \ connected_component_set (- S) a) (\x. a + r *\<^sub>R g x)" using \continuous_on (S \ C) g\ ceq by (intro continuous_intros) blast have cont2: "continuous_on (cball a r - connected_component_set (- S) a) - (\x. a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" + (\x. a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))" by (rule continuous_intros | force simp: \a \ S\)+ have 1: "continuous_on (cball a r) (\x. if connected_component (- S) a x then a + r *\<^sub>R g x - else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" + else a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))" apply (subst cb_eq) apply (rule continuous_on_cases [OF _ _ cont1 cont2]) using \closed S\ ceq cin by (force simp: closed_Diff open_Compl closed_Un_complement_component open_connected_component)+ have 2: "(\x. a + r *\<^sub>R g x) ` (cball a r \ connected_component_set (- S) a) \ sphere a r " using \0 < r\ by (force simp: dist_norm ceq) have "retraction (cball a r) (sphere a r) (\x. if x \ connected_component_set (- S) a then a + r *\<^sub>R g x - else a + r *\<^sub>R ((x - a) /\<^sub>R norm (x - a)))" + else a + r *\<^sub>R ((x-a) /\<^sub>R norm (x-a)))" using \0 < r\ \a \ S\ \a \ C\ r by (auto simp: norm_minus_commute retraction_def Pi_iff ceq dist_norm abs_if mult_less_0_iff divide_simps 1 2) then have False using no_retraction_cball [OF \0 < r\, of a, unfolded retract_of_def, simplified, rule_format, of "\x. if x \ connected_component_set (- S) a then a + r *\<^sub>R g x - else a + r *\<^sub>R inverse(norm(x - a)) *\<^sub>R (x - a)"] + else a + r *\<^sub>R inverse(norm(x-a)) *\<^sub>R (x-a)"] by blast } then show ?thesis by blast qed subsubsection \Proving surjectivity via Brouwer fixpoint theorem\ lemma brouwer_surjective: fixes f :: "'n::euclidean_space \ 'n" - assumes "compact T" - and "convex T" - and "T \ {}" - and "continuous_on T f" + assumes T: "compact T" "convex T" "T \ {}" + and f: "continuous_on T f" and "\x y. \x\S; y\T\ \ x + (y - f y) \ T" and "x \ S" shows "\y\T. f y = x" proof - have *: "\x y. f y = x \ x + (y - f y) = y" by (auto simp add: algebra_simps) show ?thesis unfolding * - apply (rule brouwer[OF assms(1-3), of "\y. x + (y - f y)"]) - apply (intro continuous_intros) - using assms - apply auto - done + proof (rule brouwer[OF T]) + show "continuous_on T (\y. x + (y - f y))" + by (intro continuous_intros f) + qed (use assms in auto) qed lemma brouwer_surjective_cball: fixes f :: "'n::euclidean_space \ 'n" assumes "continuous_on (cball a e) f" and "e > 0" and "x \ S" and "\x y. \x\S; y\cball a e\ \ x + (y - f y) \ cball a e" shows "\y\cball a e. f y = x" by (smt (verit, best) assms brouwer_surjective cball_eq_empty compact_cball convex_cball) subsubsection \Inverse function theorem\ text \See Sussmann: "Multidifferential calculus", Theorem 2.1.1\ lemma sussmann_open_mapping: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "open S" and contf: "continuous_on S f" and "x \ S" and derf: "(f has_derivative f') (at x)" and "bounded_linear g'" "f' \ g' = id" and "T \ S" and x: "x \ interior T" shows "f x \ interior (f ` T)" proof - interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto interpret g': bounded_linear g' using assms by auto obtain B where B: "0 < B" "\x. norm (g' x) \ norm x * B" using bounded_linear.pos_bounded[OF assms(5)] by blast hence *: "1 / (2 * B) > 0" by auto obtain e0 where e0: "0 < e0" "\y. norm (y - x) < e0 \ norm (f y - f x - f' (y - x)) \ 1 / (2 * B) * norm (y - x)" using derf unfolding has_derivative_at_alt using * by blast obtain e1 where e1: "0 < e1" "cball x e1 \ T" using mem_interior_cball x by blast have *: "0 < e0 / B" "0 < e1 / B" using e0 e1 B by auto obtain e where e: "0 < e" "e < e0 / B" "e < e1 / B" using field_lbound_gt_zero[OF *] by blast have lem: "\y\cball (f x) e. f (x + g' (y - f x)) = z" if "z\cball (f x) (e / 2)" for z proof (rule brouwer_surjective_cball) have z: "z \ S" if as: "y \cball (f x) e" "z = x + (g' y - g' (f x))" for y z proof- have "dist x z = norm (g' (f x) - g' y)" unfolding as(2) and dist_norm by auto also have "\ \ norm (f x - y) * B" by (metis B(2) g'.diff) also have "\ \ e * B" by (metis B(1) dist_norm mem_cball mult_le_cancel_right_pos that(1)) also have "\ \ e1" using B(1) e(3) pos_less_divide_eq by fastforce finally have "z \ cball x e1" by force then show "z \ S" using e1 assms(7) by auto qed show "continuous_on (cball (f x) e) (\y. f (x + g' (y - f x)))" unfolding g'.diff proof (rule continuous_on_compose2 [OF _ _ order_refl, of _ _ f]) show "continuous_on ((\y. x + (g' y - g' (f x))) ` cball (f x) e) f" by (rule continuous_on_subset[OF contf]) (use z in blast) show "continuous_on (cball (f x) e) (\y. x + (g' y - g' (f x)))" by (intro continuous_intros linear_continuous_on[OF \bounded_linear g'\]) qed next fix y z assume y: "y \ cball (f x) (e / 2)" and z: "z \ cball (f x) e" have "norm (g' (z - f x)) \ norm (z - f x) * B" using B by auto also have "\ \ e * B" by (metis B(1) z dist_norm mem_cball norm_minus_commute mult_le_cancel_right_pos) also have "\ < e0" using B(1) e(2) pos_less_divide_eq by blast finally have *: "norm (x + g' (z - f x) - x) < e0" by auto have **: "f x + f' (x + g' (z - f x) - x) = z" using assms(6)[unfolded o_def id_def,THEN cong] by auto have "norm (f x - (y + (z - f (x + g' (z - f x))))) \ norm (f (x + g' (z - f x)) - z) + norm (f x - y)" using norm_triangle_ineq[of "f (x + g'(z - f x)) - z" "f x - y"] by (auto simp add: algebra_simps) also have "\ \ 1 / (B * 2) * norm (g' (z - f x)) + norm (f x - y)" using e0(2)[rule_format, OF *] by (simp only: algebra_simps **) auto also have "\ \ 1 / (B * 2) * norm (g' (z - f x)) + e/2" using y by (auto simp: dist_norm) also have "\ \ 1 / (B * 2) * B * norm (z - f x) + e/2" using * B by (auto simp add: field_simps) also have "\ \ 1 / 2 * norm (z - f x) + e/2" by auto also have "\ \ e/2 + e/2" using B(1) \norm (z - f x) * B \ e * B\ by auto finally show "y + (z - f (x + g' (z - f x))) \ cball (f x) e" by (auto simp: dist_norm) qed (use e that in auto) show ?thesis unfolding mem_interior proof (intro exI conjI subsetI) fix y assume "y \ ball (f x) (e / 2)" then have *: "y \ cball (f x) (e / 2)" by auto obtain z where z: "z \ cball (f x) e" "f (x + g' (z - f x)) = y" using lem * by blast then have "norm (g' (z - f x)) \ norm (z - f x) * B" using B by (auto simp add: field_simps) also have "\ \ e * B" by (metis B(1) dist_norm mem_cball norm_minus_commute mult_le_cancel_right_pos z(1)) also have "\ \ e1" using e B unfolding less_divide_eq by auto finally have "x + g'(z - f x) \ T" by (metis add_diff_cancel diff_diff_add dist_norm e1(2) mem_cball norm_minus_commute subset_eq) then show "y \ f ` T" using z by auto qed (use e in auto) qed text \Hence the following eccentric variant of the inverse function theorem. This has no continuity assumptions, but we do need the inverse function. We could put \f' \ g = I\ but this happens to fit with the minimal linear algebra theory I've set up so far.\ lemma has_derivative_inverse_strong: fixes f :: "'n::euclidean_space \ 'n" assumes S: "open S" "x \ S" and contf: "continuous_on S f" and gf: "\x. x \ S \ g (f x) = x" and derf: "(f has_derivative f') (at x)" and id: "f' \ g' = id" shows "(g has_derivative g') (at (f x))" proof - have linf: "bounded_linear f'" using derf unfolding has_derivative_def by auto then have ling: "bounded_linear g'" unfolding linear_conv_bounded_linear[symmetric] using id right_inverse_linear by blast moreover have "g' \ f' = id" using id linear_inverse_left linear_linear linf ling by blast moreover have *: "\T. \T \ S; x \ interior T\ \ f x \ interior (f ` T)" using S derf contf id ling sussmann_open_mapping by blast have "continuous (at (f x)) g" unfolding continuous_at Lim_at proof (intro strip) fix e :: real assume "e > 0" then have "f x \ interior (f ` (ball x e \ S))" by (simp add: "*" S interior_open) then obtain d where d: "0 < d" "ball (f x) d \ f ` (ball x e \ S)" unfolding mem_interior by blast show "\d>0. \y. 0 < dist y (f x) \ dist y (f x) < d \ dist (g y) (g (f x)) < e" proof (intro exI allI impI conjI) fix y assume "0 < dist y (f x) \ dist y (f x) < d" then have "g y \ g ` f ` (ball x e \ S)" by (metis d(2) dist_commute mem_ball rev_image_eqI subset_iff) then show "dist (g y) (g (f x)) < e" using \x \ S\ by (simp add: gf dist_commute image_iff) qed (use d in auto) qed moreover have "f x \ interior (f ` S)" using "*" S interior_eq by blast moreover have "f (g y) = y" if "y \ interior (f ` S)" for y by (metis gf imageE interiorE subsetD that) ultimately show ?thesis using assms by (metis has_derivative_inverse_basic_x open_interior) qed text \A rewrite based on the other domain.\ lemma has_derivative_inverse_strong_x: fixes f :: "'a::euclidean_space \ 'a" assumes "open S" and "g y \ S" and "continuous_on S f" and "\x. x \ S \ g (f x) = x" and "(f has_derivative f') (at (g y))" and "f' \ g' = id" and f: "f (g y) = y" shows "(g has_derivative g') (at y)" using has_derivative_inverse_strong[OF assms(1-6)] by (simp add: f) text \On a region.\ theorem has_derivative_inverse_on: fixes f :: "'n::euclidean_space \ 'n" assumes "open S" and "\x. x \ S \ (f has_derivative f'(x)) (at x)" and "\x. x \ S \ g (f x) = x" and "f' x \ g' x = id" and "x \ S" shows "(g has_derivative g'(x)) (at (f x))" by (meson assms continuous_on_eq_continuous_at has_derivative_continuous has_derivative_inverse_strong) end diff --git a/src/HOL/Library/Groups_Big_Fun.thy b/src/HOL/Library/Groups_Big_Fun.thy --- a/src/HOL/Library/Groups_Big_Fun.thy +++ b/src/HOL/Library/Groups_Big_Fun.thy @@ -1,326 +1,290 @@ (* Author: Florian Haftmann, TU Muenchen *) section \Big sum and product over function bodies\ theory Groups_Big_Fun imports Main begin subsection \Abstract product\ locale comm_monoid_fun = comm_monoid begin definition G :: "('b \ 'a) \ 'a" where expand_set: "G g = comm_monoid_set.F f \<^bold>1 g {a. g a \ \<^bold>1}" interpretation F: comm_monoid_set f "\<^bold>1" .. lemma expand_superset: assumes "finite A" and "{a. g a \ \<^bold>1} \ A" shows "G g = F.F g A" - apply (simp add: expand_set) - apply (rule F.same_carrierI [of A]) - apply (simp_all add: assms) - done + using F.mono_neutral_right assms expand_set by fastforce lemma conditionalize: assumes "finite A" shows "F.F g A = G (\a. if a \ A then g a else \<^bold>1)" using assms - apply (simp add: expand_set) - apply (rule F.same_carrierI [of A]) - apply auto - done + by (smt (verit, ccfv_threshold) Diff_iff F.mono_neutral_cong_right expand_set mem_Collect_eq subsetI) + lemma neutral [simp]: "G (\a. \<^bold>1) = \<^bold>1" by (simp add: expand_set) lemma update [simp]: assumes "finite {a. g a \ \<^bold>1}" assumes "g a = \<^bold>1" shows "G (g(a := b)) = b \<^bold>* G g" proof (cases "b = \<^bold>1") case True with \g a = \<^bold>1\ show ?thesis by (simp add: expand_set) (rule F.cong, auto) next case False moreover have "{a'. a' \ a \ g a' \ \<^bold>1} = insert a {a. g a \ \<^bold>1}" by auto moreover from \g a = \<^bold>1\ have "a \ {a. g a \ \<^bold>1}" by simp moreover have "F.F (\a'. if a' = a then b else g a') {a. g a \ \<^bold>1} = F.F g {a. g a \ \<^bold>1}" by (rule F.cong) (auto simp add: \g a = \<^bold>1\) ultimately show ?thesis using \finite {a. g a \ \<^bold>1}\ by (simp add: expand_set) qed lemma infinite [simp]: "\ finite {a. g a \ \<^bold>1} \ G g = \<^bold>1" by (simp add: expand_set) lemma cong [cong]: assumes "\a. g a = h a" shows "G g = G h" using assms by (simp add: expand_set) lemma not_neutral_obtains_not_neutral: assumes "G g \ \<^bold>1" obtains a where "g a \ \<^bold>1" using assms by (auto elim: F.not_neutral_contains_not_neutral simp add: expand_set) lemma reindex_cong: assumes "bij l" assumes "g \ l = h" shows "G g = G h" proof - from assms have unfold: "h = g \ l" by simp from \bij l\ have "inj l" by (rule bij_is_inj) then have "inj_on l {a. h a \ \<^bold>1}" by (rule subset_inj_on) simp moreover from \bij l\ have "{a. g a \ \<^bold>1} = l ` {a. h a \ \<^bold>1}" by (auto simp add: image_Collect unfold elim: bij_pointE) moreover have "\x. x \ {a. h a \ \<^bold>1} \ g (l x) = h x" by (simp add: unfold) ultimately have "F.F g {a. g a \ \<^bold>1} = F.F h {a. h a \ \<^bold>1}" by (rule F.reindex_cong) then show ?thesis by (simp add: expand_set) qed lemma distrib: assumes "finite {a. g a \ \<^bold>1}" and "finite {a. h a \ \<^bold>1}" shows "G (\a. g a \<^bold>* h a) = G g \<^bold>* G h" proof - from assms have "finite ({a. g a \ \<^bold>1} \ {a. h a \ \<^bold>1})" by simp moreover have "{a. g a \<^bold>* h a \ \<^bold>1} \ {a. g a \ \<^bold>1} \ {a. h a \ \<^bold>1}" by auto (drule sym, simp) ultimately show ?thesis using assms by (simp add: expand_superset [of "{a. g a \ \<^bold>1} \ {a. h a \ \<^bold>1}"] F.distrib) qed lemma swap: assumes "finite C" assumes subset: "{a. \b. g a b \ \<^bold>1} \ {b. \a. g a b \ \<^bold>1} \ C" (is "?A \ ?B \ C") shows "G (\a. G (g a)) = G (\b. G (\a. g a b))" proof - from \finite C\ subset have "finite ({a. \b. g a b \ \<^bold>1} \ {b. \a. g a b \ \<^bold>1})" by (rule rev_finite_subset) then have fins: "finite {b. \a. g a b \ \<^bold>1}" "finite {a. \b. g a b \ \<^bold>1}" by (auto simp add: finite_cartesian_product_iff) have subsets: "\a. {b. g a b \ \<^bold>1} \ {b. \a. g a b \ \<^bold>1}" "\b. {a. g a b \ \<^bold>1} \ {a. \b. g a b \ \<^bold>1}" "{a. F.F (g a) {b. \a. g a b \ \<^bold>1} \ \<^bold>1} \ {a. \b. g a b \ \<^bold>1}" "{a. F.F (\aa. g aa a) {a. \b. g a b \ \<^bold>1} \ \<^bold>1} \ {b. \a. g a b \ \<^bold>1}" by (auto elim: F.not_neutral_contains_not_neutral) from F.swap have "F.F (\a. F.F (g a) {b. \a. g a b \ \<^bold>1}) {a. \b. g a b \ \<^bold>1} = F.F (\b. F.F (\a. g a b) {a. \b. g a b \ \<^bold>1}) {b. \a. g a b \ \<^bold>1}" . with subsets fins have "G (\a. F.F (g a) {b. \a. g a b \ \<^bold>1}) = G (\b. F.F (\a. g a b) {a. \b. g a b \ \<^bold>1})" by (auto simp add: expand_superset [of "{b. \a. g a b \ \<^bold>1}"] expand_superset [of "{a. \b. g a b \ \<^bold>1}"]) with subsets fins show ?thesis by (auto simp add: expand_superset [of "{b. \a. g a b \ \<^bold>1}"] expand_superset [of "{a. \b. g a b \ \<^bold>1}"]) qed lemma cartesian_product: assumes "finite C" assumes subset: "{a. \b. g a b \ \<^bold>1} \ {b. \a. g a b \ \<^bold>1} \ C" (is "?A \ ?B \ C") shows "G (\a. G (g a)) = G (\(a, b). g a b)" proof - from subset \finite C\ have fin_prod: "finite (?A \ ?B)" by (rule finite_subset) from fin_prod have "finite ?A" and "finite ?B" by (auto simp add: finite_cartesian_product_iff) have *: "G (\a. G (g a)) = (F.F (\a. F.F (g a) {b. \a. g a b \ \<^bold>1}) {a. \b. g a b \ \<^bold>1})" - apply (subst expand_superset [of "?B"]) - apply (rule \finite ?B\) - apply auto - apply (subst expand_superset [of "?A"]) - apply (rule \finite ?A\) - apply auto - apply (erule F.not_neutral_contains_not_neutral) - apply auto - done - have "{p. (case p of (a, b) \ g a b) \ \<^bold>1} \ ?A \ ?B" + using \finite ?A\ \finite ?B\ expand_superset + by (smt (verit, del_insts) Collect_mono local.cong not_neutral_obtains_not_neutral) + have **: "{p. (case p of (a, b) \ g a b) \ \<^bold>1} \ ?A \ ?B" by auto - with subset have **: "{p. (case p of (a, b) \ g a b) \ \<^bold>1} \ C" - by blast show ?thesis - apply (simp add: *) - apply (simp add: F.cartesian_product) - apply (subst expand_superset [of C]) - apply (rule \finite C\) - apply (simp_all add: **) - apply (rule F.same_carrierI [of C]) - apply (rule \finite C\) - apply (simp_all add: subset) - apply auto - done + using \finite C\ expand_superset + using "*" ** F.cartesian_product fin_prod by force qed lemma cartesian_product2: assumes fin: "finite D" assumes subset: "{(a, b). \c. g a b c \ \<^bold>1} \ {c. \a b. g a b c \ \<^bold>1} \ D" (is "?AB \ ?C \ D") shows "G (\(a, b). G (g a b)) = G (\(a, b, c). g a b c)" proof - have bij: "bij (\(a, b, c). ((a, b), c))" by (auto intro!: bijI injI simp add: image_def) have "{p. \c. g (fst p) (snd p) c \ \<^bold>1} \ {c. \p. g (fst p) (snd p) c \ \<^bold>1} \ D" by auto (insert subset, blast) with fin have "G (\p. G (g (fst p) (snd p))) = G (\(p, c). g (fst p) (snd p) c)" by (rule cartesian_product) then have "G (\(a, b). G (g a b)) = G (\((a, b), c). g a b c)" by (auto simp add: split_def) also have "G (\((a, b), c). g a b c) = G (\(a, b, c). g a b c)" using bij by (rule reindex_cong [of "\(a, b, c). ((a, b), c)"]) (simp add: fun_eq_iff) finally show ?thesis . qed lemma delta [simp]: "G (\b. if b = a then g b else \<^bold>1) = g a" proof - have "{b. (if b = a then g b else \<^bold>1) \ \<^bold>1} \ {a}" by auto then show ?thesis by (simp add: expand_superset [of "{a}"]) qed lemma delta' [simp]: "G (\b. if a = b then g b else \<^bold>1) = g a" proof - have "(\b. if a = b then g b else \<^bold>1) = (\b. if b = a then g b else \<^bold>1)" by (simp add: fun_eq_iff) then have "G (\b. if a = b then g b else \<^bold>1) = G (\b. if b = a then g b else \<^bold>1)" by (simp cong del: cong) then show ?thesis by simp qed end subsection \Concrete sum\ context comm_monoid_add begin sublocale Sum_any: comm_monoid_fun plus 0 rewrites "comm_monoid_set.F plus 0 = sum" defines Sum_any = Sum_any.G proof - show "comm_monoid_fun plus 0" .. then interpret Sum_any: comm_monoid_fun plus 0 . from sum_def show "comm_monoid_set.F plus 0 = sum" by (auto intro: sym) qed end syntax (ASCII) "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3SUM _. _)" [0, 10] 10) syntax "_Sum_any" :: "pttrn \ 'a \ 'a::comm_monoid_add" ("(3\_. _)" [0, 10] 10) translations "\a. b" \ "CONST Sum_any (\a. b)" lemma Sum_any_left_distrib: fixes r :: "'a :: semiring_0" assumes "finite {a. g a \ 0}" shows "Sum_any g * r = (\n. g n * r)" -proof - - note assms - moreover have "{a. g a * r \ 0} \ {a. g a \ 0}" by auto - ultimately show ?thesis - by (simp add: sum_distrib_right Sum_any.expand_superset [of "{a. g a \ 0}"]) -qed + by (metis (mono_tags, lifting) Collect_mono Sum_any.expand_superset assms mult_zero_left sum_distrib_right) lemma Sum_any_right_distrib: fixes r :: "'a :: semiring_0" assumes "finite {a. g a \ 0}" shows "r * Sum_any g = (\n. r * g n)" -proof - - note assms - moreover have "{a. r * g a \ 0} \ {a. g a \ 0}" by auto - ultimately show ?thesis - by (simp add: sum_distrib_left Sum_any.expand_superset [of "{a. g a \ 0}"]) -qed + by (metis (mono_tags, lifting) Collect_mono Sum_any.expand_superset assms mult_zero_right sum_distrib_left) lemma Sum_any_product: fixes f g :: "'b \ 'a::semiring_0" assumes "finite {a. f a \ 0}" and "finite {b. g b \ 0}" shows "Sum_any f * Sum_any g = (\a. \b. f a * g b)" proof - - have subset_f: "{a. (\b. f a * g b) \ 0} \ {a. f a \ 0}" - by rule (simp, rule, auto) - moreover have subset_g: "\a. {b. f a * g b \ 0} \ {b. g b \ 0}" - by rule (simp, rule, auto) - ultimately show ?thesis using assms - by (auto simp add: Sum_any.expand_set [of f] Sum_any.expand_set [of g] - Sum_any.expand_superset [of "{a. f a \ 0}"] Sum_any.expand_superset [of "{b. g b \ 0}"] - sum_product) + have "\a. (\b. a * g b) = a * Sum_any g" + by (simp add: Sum_any_right_distrib assms(2)) + then show ?thesis + by (simp add: Sum_any_left_distrib assms(1)) qed lemma Sum_any_eq_zero_iff [simp]: fixes f :: "'a \ nat" assumes "finite {a. f a \ 0}" shows "Sum_any f = 0 \ f = (\_. 0)" using assms by (simp add: Sum_any.expand_set fun_eq_iff) subsection \Concrete product\ context comm_monoid_mult begin sublocale Prod_any: comm_monoid_fun times 1 rewrites "comm_monoid_set.F times 1 = prod" defines Prod_any = Prod_any.G proof - show "comm_monoid_fun times 1" .. then interpret Prod_any: comm_monoid_fun times 1 . from prod_def show "comm_monoid_set.F times 1 = prod" by (auto intro: sym) qed end syntax (ASCII) "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3PROD _. _)" [0, 10] 10) syntax "_Prod_any" :: "pttrn \ 'a \ 'a::comm_monoid_mult" ("(3\_. _)" [0, 10] 10) translations "\a. b" == "CONST Prod_any (\a. b)" lemma Prod_any_zero: fixes f :: "'b \ 'a :: comm_semiring_1" assumes "finite {a. f a \ 1}" assumes "f a = 0" shows "(\a. f a) = 0" proof - from \f a = 0\ have "f a \ 1" by simp with \f a = 0\ have "\a. f a \ 1 \ f a = 0" by blast with \finite {a. f a \ 1}\ show ?thesis by (simp add: Prod_any.expand_set prod_zero) qed lemma Prod_any_not_zero: fixes f :: "'b \ 'a :: comm_semiring_1" assumes "finite {a. f a \ 1}" assumes "(\a. f a) \ 0" shows "f a \ 0" using assms Prod_any_zero [of f] by blast lemma power_Sum_any: assumes "finite {a. f a \ 0}" shows "c ^ (\a. f a) = (\a. c ^ f a)" proof - have "{a. c ^ f a \ 1} \ {a. f a \ 0}" by (auto intro: ccontr) with assms show ?thesis by (simp add: Sum_any.expand_set Prod_any.expand_superset power_sum) qed end diff --git a/src/HOL/Library/Multiset.thy b/src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy +++ b/src/HOL/Library/Multiset.thy @@ -1,4782 +1,4677 @@ (* Title: HOL/Library/Multiset.thy Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker Author: Andrei Popescu, TU Muenchen Author: Jasmin Blanchette, Inria, LORIA, MPII Author: Dmitriy Traytel, TU Muenchen Author: Mathias Fleury, MPII Author: Martin Desharnais, MPI-INF Saarbruecken *) section \(Finite) Multisets\ theory Multiset imports Cancellation begin subsection \The type of multisets\ typedef 'a multiset = \{f :: 'a \ nat. finite {x. f x > 0}}\ morphisms count Abs_multiset proof show \(\x. 0::nat) \ {f. finite {x. f x > 0}}\ by simp qed setup_lifting type_definition_multiset lemma count_Abs_multiset: \count (Abs_multiset f) = f\ if \finite {x. f x > 0}\ by (rule Abs_multiset_inverse) (simp add: that) lemma multiset_eq_iff: "M = N \ (\a. count M a = count N a)" by (simp only: count_inject [symmetric] fun_eq_iff) lemma multiset_eqI: "(\x. count A x = count B x) \ A = B" using multiset_eq_iff by auto text \Preservation of the representing set \<^term>\multiset\.\ lemma diff_preserves_multiset: \finite {x. 0 < M x - N x}\ if \finite {x. 0 < M x}\ for M N :: \'a \ nat\ using that by (rule rev_finite_subset) auto lemma filter_preserves_multiset: \finite {x. 0 < (if P x then M x else 0)}\ if \finite {x. 0 < M x}\ for M N :: \'a \ nat\ using that by (rule rev_finite_subset) auto lemmas in_multiset = diff_preserves_multiset filter_preserves_multiset subsection \Representing multisets\ text \Multiset enumeration\ instantiation multiset :: (type) cancel_comm_monoid_add begin lift_definition zero_multiset :: \'a multiset\ is \\a. 0\ by simp abbreviation empty_mset :: \'a multiset\ (\{#}\) where \empty_mset \ 0\ lift_definition plus_multiset :: \'a multiset \ 'a multiset \ 'a multiset\ is \\M N a. M a + N a\ by simp lift_definition minus_multiset :: \'a multiset \ 'a multiset \ 'a multiset\ is \\M N a. M a - N a\ by (rule diff_preserves_multiset) instance by (standard; transfer) (simp_all add: fun_eq_iff) end context begin qualified definition is_empty :: "'a multiset \ bool" where [code_abbrev]: "is_empty A \ A = {#}" end lemma add_mset_in_multiset: \finite {x. 0 < (if x = a then Suc (M x) else M x)}\ if \finite {x. 0 < M x}\ using that by (simp add: flip: insert_Collect) lift_definition add_mset :: "'a \ 'a multiset \ 'a multiset" is "\a M b. if b = a then Suc (M b) else M b" by (rule add_mset_in_multiset) syntax "_multiset" :: "args \ 'a multiset" ("{#(_)#}") translations "{#x, xs#}" == "CONST add_mset x {#xs#}" "{#x#}" == "CONST add_mset x {#}" lemma count_empty [simp]: "count {#} a = 0" by (simp add: zero_multiset.rep_eq) lemma count_add_mset [simp]: "count (add_mset b A) a = (if b = a then Suc (count A a) else count A a)" by (simp add: add_mset.rep_eq) lemma count_single: "count {#b#} a = (if b = a then 1 else 0)" by simp lemma add_mset_not_empty [simp]: \add_mset a A \ {#}\ and empty_not_add_mset [simp]: "{#} \ add_mset a A" by (auto simp: multiset_eq_iff) lemma add_mset_add_mset_same_iff [simp]: "add_mset a A = add_mset a B \ A = B" by (auto simp: multiset_eq_iff) lemma add_mset_commute: "add_mset x (add_mset y M) = add_mset y (add_mset x M)" by (auto simp: multiset_eq_iff) subsection \Basic operations\ subsubsection \Conversion to set and membership\ definition set_mset :: \'a multiset \ 'a set\ where \set_mset M = {x. count M x > 0}\ abbreviation member_mset :: \'a \ 'a multiset \ bool\ where \member_mset a M \ a \ set_mset M\ notation member_mset (\'(\#')\) and member_mset (\(_/ \# _)\ [50, 51] 50) notation (ASCII) member_mset (\'(:#')\) and member_mset (\(_/ :# _)\ [50, 51] 50) abbreviation not_member_mset :: \'a \ 'a multiset \ bool\ where \not_member_mset a M \ a \ set_mset M\ notation not_member_mset (\'(\#')\) and not_member_mset (\(_/ \# _)\ [50, 51] 50) notation (ASCII) not_member_mset (\'(~:#')\) and not_member_mset (\(_/ ~:# _)\ [50, 51] 50) context begin qualified abbreviation Ball :: "'a multiset \ ('a \ bool) \ bool" where "Ball M \ Set.Ball (set_mset M)" qualified abbreviation Bex :: "'a multiset \ ('a \ bool) \ bool" where "Bex M \ Set.Bex (set_mset M)" end syntax "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_\#_./ _)" [0, 0, 10] 10) syntax (ASCII) "_MBall" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) "_MBex" :: "pttrn \ 'a set \ bool \ bool" ("(3\_:#_./ _)" [0, 0, 10] 10) translations "\x\#A. P" \ "CONST Multiset.Ball A (\x. P)" "\x\#A. P" \ "CONST Multiset.Bex A (\x. P)" print_translation \ [Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Multiset.Ball\ \<^syntax_const>\_MBall\, Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\Multiset.Bex\ \<^syntax_const>\_MBex\] \ \ \to avoid eta-contraction of body\ lemma count_eq_zero_iff: "count M x = 0 \ x \# M" by (auto simp add: set_mset_def) lemma not_in_iff: "x \# M \ count M x = 0" by (auto simp add: count_eq_zero_iff) lemma count_greater_zero_iff [simp]: "count M x > 0 \ x \# M" by (auto simp add: set_mset_def) lemma count_inI: assumes "count M x = 0 \ False" shows "x \# M" proof (rule ccontr) assume "x \# M" with assms show False by (simp add: not_in_iff) qed lemma in_countE: assumes "x \# M" obtains n where "count M x = Suc n" proof - from assms have "count M x > 0" by simp then obtain n where "count M x = Suc n" using gr0_conv_Suc by blast with that show thesis . qed lemma count_greater_eq_Suc_zero_iff [simp]: "count M x \ Suc 0 \ x \# M" by (simp add: Suc_le_eq) lemma count_greater_eq_one_iff [simp]: "count M x \ 1 \ x \# M" by simp lemma set_mset_empty [simp]: "set_mset {#} = {}" by (simp add: set_mset_def) lemma set_mset_single: "set_mset {#b#} = {b}" by (simp add: set_mset_def) lemma set_mset_eq_empty_iff [simp]: "set_mset M = {} \ M = {#}" by (auto simp add: multiset_eq_iff count_eq_zero_iff) lemma finite_set_mset [iff]: "finite (set_mset M)" using count [of M] by simp lemma set_mset_add_mset_insert [simp]: \set_mset (add_mset a A) = insert a (set_mset A)\ by (auto simp flip: count_greater_eq_Suc_zero_iff split: if_splits) lemma multiset_nonemptyE [elim]: assumes "A \ {#}" obtains x where "x \# A" proof - have "\x. x \# A" by (rule ccontr) (insert assms, auto) with that show ?thesis by blast qed lemma count_gt_imp_in_mset: "count M x > n \ x \# M" using count_greater_zero_iff by fastforce subsubsection \Union\ lemma count_union [simp]: "count (M + N) a = count M a + count N a" by (simp add: plus_multiset.rep_eq) lemma set_mset_union [simp]: "set_mset (M + N) = set_mset M \ set_mset N" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_union) simp lemma union_mset_add_mset_left [simp]: "add_mset a A + B = add_mset a (A + B)" by (auto simp: multiset_eq_iff) lemma union_mset_add_mset_right [simp]: "A + add_mset a B = add_mset a (A + B)" by (auto simp: multiset_eq_iff) lemma add_mset_add_single: \add_mset a A = A + {#a#}\ by (subst union_mset_add_mset_right, subst add.comm_neutral) standard subsubsection \Difference\ instance multiset :: (type) comm_monoid_diff by standard (transfer; simp add: fun_eq_iff) lemma count_diff [simp]: "count (M - N) a = count M a - count N a" by (simp add: minus_multiset.rep_eq) lemma add_mset_diff_bothsides: \add_mset a M - add_mset a A = M - A\ by (auto simp: multiset_eq_iff) lemma in_diff_count: "a \# M - N \ count N a < count M a" by (simp add: set_mset_def) lemma count_in_diffI: assumes "\n. count N x = n + count M x \ False" shows "x \# M - N" proof (rule ccontr) assume "x \# M - N" then have "count N x = (count N x - count M x) + count M x" by (simp add: in_diff_count not_less) with assms show False by auto qed lemma in_diff_countE: assumes "x \# M - N" obtains n where "count M x = Suc n + count N x" proof - from assms have "count M x - count N x > 0" by (simp add: in_diff_count) then have "count M x > count N x" by simp then obtain n where "count M x = Suc n + count N x" using less_iff_Suc_add by auto with that show thesis . qed lemma in_diffD: assumes "a \# M - N" shows "a \# M" proof - have "0 \ count N a" by simp also from assms have "count N a < count M a" by (simp add: in_diff_count) finally show ?thesis by simp qed lemma set_mset_diff: "set_mset (M - N) = {a. count N a < count M a}" by (simp add: set_mset_def) lemma diff_empty [simp]: "M - {#} = M \ {#} - M = {#}" by rule (fact Groups.diff_zero, fact Groups.zero_diff) lemma diff_cancel: "A - A = {#}" by (fact Groups.diff_cancel) lemma diff_union_cancelR: "M + N - N = (M::'a multiset)" by (fact add_diff_cancel_right') lemma diff_union_cancelL: "N + M - N = (M::'a multiset)" by (fact add_diff_cancel_left') lemma diff_right_commute: fixes M N Q :: "'a multiset" shows "M - N - Q = M - Q - N" by (fact diff_right_commute) lemma diff_add: fixes M N Q :: "'a multiset" shows "M - (N + Q) = M - N - Q" by (rule sym) (fact diff_diff_add) lemma insert_DiffM [simp]: "x \# M \ add_mset x (M - {#x#}) = M" by (clarsimp simp: multiset_eq_iff) lemma insert_DiffM2: "x \# M \ (M - {#x#}) + {#x#} = M" by simp lemma diff_union_swap: "a \ b \ add_mset b (M - {#a#}) = add_mset b M - {#a#}" by (auto simp add: multiset_eq_iff) lemma diff_add_mset_swap [simp]: "b \# A \ add_mset b M - A = add_mset b (M - A)" by (auto simp add: multiset_eq_iff simp: not_in_iff) lemma diff_union_swap2 [simp]: "y \# M \ add_mset x M - {#y#} = add_mset x (M - {#y#})" by (metis add_mset_diff_bothsides diff_union_swap diff_zero insert_DiffM) lemma diff_diff_add_mset [simp]: "(M::'a multiset) - N - P = M - (N + P)" by (rule diff_diff_add) lemma diff_union_single_conv: "a \# J \ I + J - {#a#} = I + (J - {#a#})" by (simp add: multiset_eq_iff Suc_le_eq) lemma mset_add [elim?]: assumes "a \# A" obtains B where "A = add_mset a B" proof - from assms have "A = add_mset a (A - {#a#})" by simp with that show thesis . qed lemma union_iff: "a \# A + B \ a \# A \ a \# B" by auto lemma count_minus_inter_lt_count_minus_inter_iff: "count (M2 - M1) y < count (M1 - M2) y \ y \# M1 - M2" by (meson count_greater_zero_iff gr_implies_not_zero in_diff_count leI order.strict_trans2 order_less_asym) lemma minus_inter_eq_minus_inter_iff: "(M1 - M2) = (M2 - M1) \ set_mset (M1 - M2) = set_mset (M2 - M1)" by (metis add.commute count_diff count_eq_zero_iff diff_add_zero in_diff_countE multiset_eq_iff) subsubsection \Min and Max\ abbreviation Min_mset :: "'a::linorder multiset \ 'a" where "Min_mset m \ Min (set_mset m)" abbreviation Max_mset :: "'a::linorder multiset \ 'a" where "Max_mset m \ Max (set_mset m)" lemma Min_in_mset: "M \ {#} \ Min_mset M \# M" and Max_in_mset: "M \ {#} \ Max_mset M \# M" by simp+ subsubsection \Equality of multisets\ lemma single_eq_single [simp]: "{#a#} = {#b#} \ a = b" by (auto simp add: multiset_eq_iff) lemma union_eq_empty [iff]: "M + N = {#} \ M = {#} \ N = {#}" by (auto simp add: multiset_eq_iff) lemma empty_eq_union [iff]: "{#} = M + N \ M = {#} \ N = {#}" by (auto simp add: multiset_eq_iff) lemma multi_self_add_other_not_self [simp]: "M = add_mset x M \ False" by (auto simp add: multiset_eq_iff) lemma add_mset_remove_trivial [simp]: \add_mset x M - {#x#} = M\ by (auto simp: multiset_eq_iff) lemma diff_single_trivial: "\ x \# M \ M - {#x#} = M" by (auto simp add: multiset_eq_iff not_in_iff) lemma diff_single_eq_union: "x \# M \ M - {#x#} = N \ M = add_mset x N" by auto lemma union_single_eq_diff: "add_mset x M = N \ M = N - {#x#}" unfolding add_mset_add_single[of _ M] by (fact add_implies_diff) lemma union_single_eq_member: "add_mset x M = N \ x \# N" by auto lemma add_mset_remove_trivial_If: "add_mset a (N - {#a#}) = (if a \# N then N else add_mset a N)" by (simp add: diff_single_trivial) lemma add_mset_remove_trivial_eq: \N = add_mset a (N - {#a#}) \ a \# N\ by (auto simp: add_mset_remove_trivial_If) lemma union_is_single: "M + N = {#a#} \ M = {#a#} \ N = {#} \ M = {#} \ N = {#a#}" (is "?lhs = ?rhs") proof show ?lhs if ?rhs using that by auto show ?rhs if ?lhs by (metis Multiset.diff_cancel add.commute add_diff_cancel_left' diff_add_zero diff_single_trivial insert_DiffM that) qed lemma single_is_union: "{#a#} = M + N \ {#a#} = M \ N = {#} \ M = {#} \ {#a#} = N" by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single) lemma add_eq_conv_diff: "add_mset a M = add_mset b N \ M = N \ a = b \ M = add_mset b (N - {#a#}) \ N = add_mset a (M - {#b#})" (is "?lhs \ ?rhs") (* shorter: by (simp add: multiset_eq_iff) fastforce *) proof show ?lhs if ?rhs using that by (auto simp add: add_mset_commute[of a b]) show ?rhs if ?lhs proof (cases "a = b") case True with \?lhs\ show ?thesis by simp next case False from \?lhs\ have "a \# add_mset b N" by (rule union_single_eq_member) with False have "a \# N" by auto moreover from \?lhs\ have "M = add_mset b N - {#a#}" by (rule union_single_eq_diff) moreover note False ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"]) qed qed lemma add_mset_eq_single [iff]: "add_mset b M = {#a#} \ b = a \ M = {#}" by (auto simp: add_eq_conv_diff) lemma single_eq_add_mset [iff]: "{#a#} = add_mset b M \ b = a \ M = {#}" by (auto simp: add_eq_conv_diff) lemma insert_noteq_member: assumes BC: "add_mset b B = add_mset c C" and bnotc: "b \ c" shows "c \# B" proof - have "c \# add_mset c C" by simp have nc: "\ c \# {#b#}" using bnotc by simp then have "c \# add_mset b B" using BC by simp then show "c \# B" using nc by simp qed lemma add_eq_conv_ex: "(add_mset a M = add_mset b N) = (M = N \ a = b \ (\K. M = add_mset b K \ N = add_mset a K))" by (auto simp add: add_eq_conv_diff) lemma multi_member_split: "x \# M \ \A. M = add_mset x A" by (rule exI [where x = "M - {#x#}"]) simp lemma multiset_add_sub_el_shuffle: assumes "c \# B" and "b \ c" shows "add_mset b (B - {#c#}) = add_mset b B - {#c#}" proof - from \c \# B\ obtain A where B: "B = add_mset c A" by (blast dest: multi_member_split) have "add_mset b A = add_mset c (add_mset b A) - {#c#}" by simp then have "add_mset b A = add_mset b (add_mset c A) - {#c#}" by (simp add: \b \ c\) then show ?thesis using B by simp qed lemma add_mset_eq_singleton_iff[iff]: "add_mset x M = {#y#} \ M = {#} \ x = y" by auto subsubsection \Pointwise ordering induced by count\ definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "A \# B \ (\a. count A a \ count B a)" definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "A \# B \ A \# B \ A \ B" abbreviation (input) supseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "supseteq_mset A B \ B \# A" abbreviation (input) supset_mset :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where "supset_mset A B \ B \# A" notation (input) subseteq_mset (infix "\#" 50) and supseteq_mset (infix "\#" 50) notation (ASCII) subseteq_mset (infix "<=#" 50) and subset_mset (infix "<#" 50) and supseteq_mset (infix ">=#" 50) and supset_mset (infix ">#" 50) global_interpretation subset_mset: ordering \(\#)\ \(\#)\ by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order.trans order.antisym) interpretation subset_mset: ordered_ab_semigroup_add_imp_le \(+)\ \(-)\ \(\#)\ \(\#)\ by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym) \ \FIXME: avoid junk stemming from type class interpretation\ interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "(+)" 0 "(-)" "(\#)" "(\#)" by standard \ \FIXME: avoid junk stemming from type class interpretation\ lemma mset_subset_eqI: "(\a. count A a \ count B a) \ A \# B" by (simp add: subseteq_mset_def) lemma mset_subset_eq_count: "A \# B \ count A a \ count B a" by (simp add: subseteq_mset_def) lemma mset_subset_eq_exists_conv: "(A::'a multiset) \# B \ (\C. B = A + C)" unfolding subseteq_mset_def - apply (rule iffI) - apply (rule exI [where x = "B - A"]) - apply (auto intro: multiset_eq_iff [THEN iffD2]) - done + by (metis add_diff_cancel_left' count_diff count_union le_Suc_ex le_add_same_cancel1 multiset_eq_iff zero_le) interpretation subset_mset: ordered_cancel_comm_monoid_diff "(+)" 0 "(\#)" "(\#)" "(-)" by standard (simp, fact mset_subset_eq_exists_conv) \ \FIXME: avoid junk stemming from type class interpretation\ declare subset_mset.add_diff_assoc[simp] subset_mset.add_diff_assoc2[simp] lemma mset_subset_eq_mono_add_right_cancel: "(A::'a multiset) + C \# B + C \ A \# B" by (fact subset_mset.add_le_cancel_right) lemma mset_subset_eq_mono_add_left_cancel: "C + (A::'a multiset) \# C + B \ A \# B" by (fact subset_mset.add_le_cancel_left) lemma mset_subset_eq_mono_add: "(A::'a multiset) \# B \ C \# D \ A + C \# B + D" by (fact subset_mset.add_mono) lemma mset_subset_eq_add_left: "(A::'a multiset) \# A + B" by simp lemma mset_subset_eq_add_right: "B \# (A::'a multiset) + B" by simp lemma single_subset_iff [simp]: "{#a#} \# M \ a \# M" by (auto simp add: subseteq_mset_def Suc_le_eq) lemma mset_subset_eq_single: "a \# B \ {#a#} \# B" by simp lemma mset_subset_eq_add_mset_cancel: \add_mset a A \# add_mset a B \ A \# B\ unfolding add_mset_add_single[of _ A] add_mset_add_single[of _ B] by (rule mset_subset_eq_mono_add_right_cancel) lemma multiset_diff_union_assoc: fixes A B C D :: "'a multiset" shows "C \# B \ A + B - C = A + (B - C)" by (fact subset_mset.diff_add_assoc) lemma mset_subset_eq_multiset_union_diff_commute: fixes A B C D :: "'a multiset" shows "B \# A \ A - B + C = A + C - B" by (fact subset_mset.add_diff_assoc2) lemma diff_subset_eq_self[simp]: "(M::'a multiset) - N \# M" by (simp add: subseteq_mset_def) lemma mset_subset_eqD: assumes "A \# B" and "x \# A" shows "x \# B" proof - from \x \# A\ have "count A x > 0" by simp also from \A \# B\ have "count A x \ count B x" by (simp add: subseteq_mset_def) finally show ?thesis by simp qed lemma mset_subsetD: "A \# B \ x \# A \ x \# B" by (auto intro: mset_subset_eqD [of A]) lemma set_mset_mono: "A \# B \ set_mset A \ set_mset B" by (metis mset_subset_eqD subsetI) lemma mset_subset_eq_insertD: - "add_mset x A \# B \ x \# B \ A \# B" -apply (rule conjI) - apply (simp add: mset_subset_eqD) - apply (clarsimp simp: subset_mset_def subseteq_mset_def) - apply safe - apply (erule_tac x = a in allE) - apply (auto split: if_split_asm) -done + assumes "add_mset x A \# B" + shows "x \# B \ A \# B" +proof + show "x \# B" + using assms by (simp add: mset_subset_eqD) + have "A \# add_mset x A" + by (metis (no_types) add_mset_add_single mset_subset_eq_add_left) + then have "A \# add_mset x A" + by (meson multi_self_add_other_not_self subset_mset.le_imp_less_or_eq) + then show "A \# B" + using assms subset_mset.strict_trans2 by blast +qed lemma mset_subset_insertD: "add_mset x A \# B \ x \# B \ A \# B" by (rule mset_subset_eq_insertD) simp lemma mset_subset_of_empty[simp]: "A \# {#} \ False" by (simp only: subset_mset.not_less_zero) lemma empty_subset_add_mset[simp]: "{#} \# add_mset x M" by (auto intro: subset_mset.gr_zeroI) lemma empty_le: "{#} \# A" by (fact subset_mset.zero_le) lemma insert_subset_eq_iff: "add_mset a A \# B \ a \# B \ A \# B - {#a#}" - using le_diff_conv2 [of "Suc 0" "count B a" "count A a"] - apply (auto simp add: subseteq_mset_def not_in_iff Suc_le_eq) - apply (rule ccontr) - apply (auto simp add: not_in_iff) - done + using mset_subset_eq_insertD subset_mset.le_diff_conv2 by fastforce lemma insert_union_subset_iff: "add_mset a A \# B \ a \# B \ A \# B - {#a#}" by (auto simp add: insert_subset_eq_iff subset_mset_def) lemma subset_eq_diff_conv: "A - C \# B \ A \# B + C" by (simp add: subseteq_mset_def le_diff_conv) lemma multi_psub_of_add_self [simp]: "A \# add_mset x A" by (auto simp: subset_mset_def subseteq_mset_def) lemma multi_psub_self: "A \# A = False" by simp lemma mset_subset_add_mset [simp]: "add_mset x N \# add_mset x M \ N \# M" unfolding add_mset_add_single[of _ N] add_mset_add_single[of _ M] by (fact subset_mset.add_less_cancel_right) lemma mset_subset_diff_self: "c \# B \ B - {#c#} \# B" by (auto simp: subset_mset_def elim: mset_add) lemma Diff_eq_empty_iff_mset: "A - B = {#} \ A \# B" by (auto simp: multiset_eq_iff subseteq_mset_def) lemma add_mset_subseteq_single_iff[iff]: "add_mset a M \# {#b#} \ M = {#} \ a = b" proof assume A: "add_mset a M \# {#b#}" then have \a = b\ by (auto dest: mset_subset_eq_insertD) then show "M={#} \ a=b" using A by (simp add: mset_subset_eq_add_mset_cancel) qed simp lemma nonempty_subseteq_mset_eq_single: "M \ {#} \ M \# {#x#} \ M = {#x#}" by (cases M) (metis single_is_union subset_mset.less_eqE) lemma nonempty_subseteq_mset_iff_single: "(M \ {#} \ M \# {#x#} \ P) \ M = {#x#} \ P" by (cases M) (metis empty_not_add_mset nonempty_subseteq_mset_eq_single subset_mset.order_refl) subsubsection \Intersection and bounded union\ definition inter_mset :: \'a multiset \ 'a multiset \ 'a multiset\ (infixl \\#\ 70) where \A \# B = A - (A - B)\ lemma count_inter_mset [simp]: \count (A \# B) x = min (count A x) (count B x)\ by (simp add: inter_mset_def) (*global_interpretation subset_mset: semilattice_order \(\#)\ \(\#)\ \(\#)\ by standard (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def min_def)*) interpretation subset_mset: semilattice_inf \(\#)\ \(\#)\ \(\#)\ by standard (simp_all add: multiset_eq_iff subseteq_mset_def) \ \FIXME: avoid junk stemming from type class interpretation\ definition union_mset :: \'a multiset \ 'a multiset \ 'a multiset\ (infixl \\#\ 70) where \A \# B = A + (B - A)\ lemma count_union_mset [simp]: \count (A \# B) x = max (count A x) (count B x)\ by (simp add: union_mset_def) global_interpretation subset_mset: semilattice_neutr_order \(\#)\ \{#}\ \(\#)\ \(\#)\ - apply standard - apply (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def max_def) - apply (auto simp add: le_antisym dest: sym) - apply (metis nat_le_linear)+ - done +proof + show "\a b. (b \# a) = (a = a \# b)" + by (simp add: Diff_eq_empty_iff_mset union_mset_def) + show "\a b. (b \# a) = (a = a \# b \ a \ b)" + by (metis Diff_eq_empty_iff_mset add_cancel_left_right subset_mset_def union_mset_def) +qed (auto simp: multiset_eqI union_mset_def) interpretation subset_mset: semilattice_sup \(\#)\ \(\#)\ \(\#)\ proof - have [simp]: "m \ n \ q \ n \ m + (q - m) \ n" for m n q :: nat by arith show "class.semilattice_sup (\#) (\#) (\#)" by standard (auto simp add: union_mset_def subseteq_mset_def) qed \ \FIXME: avoid junk stemming from type class interpretation\ interpretation subset_mset: bounded_lattice_bot "(\#)" "(\#)" "(\#)" "(\#)" "{#}" by standard auto \ \FIXME: avoid junk stemming from type class interpretation\ subsubsection \Additional intersection facts\ lemma set_mset_inter [simp]: "set_mset (A \# B) = set_mset A \ set_mset B" by (simp only: set_mset_def) auto lemma diff_intersect_left_idem [simp]: "M - M \# N = M - N" by (simp add: multiset_eq_iff min_def) lemma diff_intersect_right_idem [simp]: "M - N \# M = M - N" by (simp add: multiset_eq_iff min_def) lemma multiset_inter_single[simp]: "a \ b \ {#a#} \# {#b#} = {#}" by (rule multiset_eqI) auto lemma multiset_union_diff_commute: assumes "B \# C = {#}" shows "A + B - C = A - C + B" proof (rule multiset_eqI) fix x from assms have "min (count B x) (count C x) = 0" by (auto simp add: multiset_eq_iff) then have "count B x = 0 \ count C x = 0" unfolding min_def by (auto split: if_splits) then show "count (A + B - C) x = count (A - C + B) x" by auto qed lemma disjunct_not_in: - "A \# B = {#} \ (\a. a \# A \ a \# B)" (is "?P \ ?Q") -proof - assume ?P - show ?Q - proof - fix a - from \?P\ have "min (count A a) (count B a) = 0" - by (simp add: multiset_eq_iff) - then have "count A a = 0 \ count B a = 0" - by (cases "count A a \ count B a") (simp_all add: min_def) - then show "a \# A \ a \# B" - by (simp add: not_in_iff) - qed -next - assume ?Q - show ?P - proof (rule multiset_eqI) - fix a - from \?Q\ have "count A a = 0 \ count B a = 0" - by (auto simp add: not_in_iff) - then show "count (A \# B) a = count {#} a" - by auto - qed -qed + "A \# B = {#} \ (\a. a \# A \ a \# B)" + by (metis disjoint_iff set_mset_eq_empty_iff set_mset_inter) lemma inter_mset_empty_distrib_right: "A \# (B + C) = {#} \ A \# B = {#} \ A \# C = {#}" by (meson disjunct_not_in union_iff) lemma inter_mset_empty_distrib_left: "(A + B) \# C = {#} \ A \# C = {#} \ B \# C = {#}" by (meson disjunct_not_in union_iff) lemma add_mset_inter_add_mset [simp]: "add_mset a A \# add_mset a B = add_mset a (A \# B)" by (rule multiset_eqI) simp lemma add_mset_disjoint [simp]: "add_mset a A \# B = {#} \ a \# B \ A \# B = {#}" "{#} = add_mset a A \# B \ a \# B \ {#} = A \# B" by (auto simp: disjunct_not_in) lemma disjoint_add_mset [simp]: "B \# add_mset a A = {#} \ a \# B \ B \# A = {#}" "{#} = A \# add_mset b B \ b \# A \ {#} = A \# B" by (auto simp: disjunct_not_in) lemma inter_add_left1: "\ x \# N \ (add_mset x M) \# N = M \# N" by (simp add: multiset_eq_iff not_in_iff) lemma inter_add_left2: "x \# N \ (add_mset x M) \# N = add_mset x (M \# (N - {#x#}))" by (auto simp add: multiset_eq_iff elim: mset_add) lemma inter_add_right1: "\ x \# N \ N \# (add_mset x M) = N \# M" by (simp add: multiset_eq_iff not_in_iff) lemma inter_add_right2: "x \# N \ N \# (add_mset x M) = add_mset x ((N - {#x#}) \# M)" by (auto simp add: multiset_eq_iff elim: mset_add) lemma disjunct_set_mset_diff: assumes "M \# N = {#}" shows "set_mset (M - N) = set_mset M" proof (rule set_eqI) fix a from assms have "a \# M \ a \# N" by (simp add: disjunct_not_in) then show "a \# M - N \ a \# M" by (auto dest: in_diffD) (simp add: in_diff_count not_in_iff) qed lemma at_most_one_mset_mset_diff: assumes "a \# M - {#a#}" shows "set_mset (M - {#a#}) = set_mset M - {a}" using assms by (auto simp add: not_in_iff in_diff_count set_eq_iff) lemma more_than_one_mset_mset_diff: assumes "a \# M - {#a#}" shows "set_mset (M - {#a#}) = set_mset M" proof (rule set_eqI) fix b have "Suc 0 < count M b \ count M b > 0" by arith then show "b \# M - {#a#} \ b \# M" using assms by (auto simp add: in_diff_count) qed lemma inter_iff: "a \# A \# B \ a \# A \ a \# B" by simp lemma inter_union_distrib_left: "A \# B + C = (A + C) \# (B + C)" by (simp add: multiset_eq_iff min_add_distrib_left) lemma inter_union_distrib_right: "C + A \# B = (C + A) \# (C + B)" using inter_union_distrib_left [of A B C] by (simp add: ac_simps) lemma inter_subset_eq_union: "A \# B \# A + B" by (auto simp add: subseteq_mset_def) subsubsection \Additional bounded union facts\ lemma set_mset_sup [simp]: \set_mset (A \# B) = set_mset A \ set_mset B\ by (simp only: set_mset_def) (auto simp add: less_max_iff_disj) lemma sup_union_left1 [simp]: "\ x \# N \ (add_mset x M) \# N = add_mset x (M \# N)" by (simp add: multiset_eq_iff not_in_iff) lemma sup_union_left2: "x \# N \ (add_mset x M) \# N = add_mset x (M \# (N - {#x#}))" by (simp add: multiset_eq_iff) lemma sup_union_right1 [simp]: "\ x \# N \ N \# (add_mset x M) = add_mset x (N \# M)" by (simp add: multiset_eq_iff not_in_iff) lemma sup_union_right2: "x \# N \ N \# (add_mset x M) = add_mset x ((N - {#x#}) \# M)" by (simp add: multiset_eq_iff) lemma sup_union_distrib_left: "A \# B + C = (A + C) \# (B + C)" by (simp add: multiset_eq_iff max_add_distrib_left) lemma union_sup_distrib_right: "C + A \# B = (C + A) \# (C + B)" using sup_union_distrib_left [of A B C] by (simp add: ac_simps) lemma union_diff_inter_eq_sup: "A + B - A \# B = A \# B" by (auto simp add: multiset_eq_iff) lemma union_diff_sup_eq_inter: "A + B - A \# B = A \# B" by (auto simp add: multiset_eq_iff) lemma add_mset_union: \add_mset a A \# add_mset a B = add_mset a (A \# B)\ by (auto simp: multiset_eq_iff max_def) subsection \Replicate and repeat operations\ definition replicate_mset :: "nat \ 'a \ 'a multiset" where "replicate_mset n x = (add_mset x ^^ n) {#}" lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}" unfolding replicate_mset_def by simp lemma replicate_mset_Suc [simp]: "replicate_mset (Suc n) x = add_mset x (replicate_mset n x)" unfolding replicate_mset_def by (induct n) (auto intro: add.commute) lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)" unfolding replicate_mset_def by (induct n) auto lift_definition repeat_mset :: \nat \ 'a multiset \ 'a multiset\ is \\n M a. n * M a\ by simp lemma count_repeat_mset [simp]: "count (repeat_mset i A) a = i * count A a" by transfer rule lemma repeat_mset_0 [simp]: \repeat_mset 0 M = {#}\ by transfer simp lemma repeat_mset_Suc [simp]: \repeat_mset (Suc n) M = M + repeat_mset n M\ by transfer simp lemma repeat_mset_right [simp]: "repeat_mset a (repeat_mset b A) = repeat_mset (a * b) A" by (auto simp: multiset_eq_iff left_diff_distrib') lemma left_diff_repeat_mset_distrib': \repeat_mset (i - j) u = repeat_mset i u - repeat_mset j u\ by (auto simp: multiset_eq_iff left_diff_distrib') lemma left_add_mult_distrib_mset: "repeat_mset i u + (repeat_mset j u + k) = repeat_mset (i+j) u + k" by (auto simp: multiset_eq_iff add_mult_distrib) lemma repeat_mset_distrib: "repeat_mset (m + n) A = repeat_mset m A + repeat_mset n A" by (auto simp: multiset_eq_iff Nat.add_mult_distrib) lemma repeat_mset_distrib2[simp]: "repeat_mset n (A + B) = repeat_mset n A + repeat_mset n B" by (auto simp: multiset_eq_iff add_mult_distrib2) lemma repeat_mset_replicate_mset[simp]: "repeat_mset n {#a#} = replicate_mset n a" by (auto simp: multiset_eq_iff) lemma repeat_mset_distrib_add_mset[simp]: "repeat_mset n (add_mset a A) = replicate_mset n a + repeat_mset n A" by (auto simp: multiset_eq_iff) lemma repeat_mset_empty[simp]: "repeat_mset n {#} = {#}" by transfer simp subsubsection \Simprocs\ lemma repeat_mset_iterate_add: \repeat_mset n M = iterate_add n M\ unfolding iterate_add_def by (induction n) auto lemma mset_subseteq_add_iff1: "j \ (i::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (repeat_mset (i-j) u + m \# n)" by (auto simp add: subseteq_mset_def nat_le_add_iff1) lemma mset_subseteq_add_iff2: "i \ (j::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (m \# repeat_mset (j-i) u + n)" by (auto simp add: subseteq_mset_def nat_le_add_iff2) lemma mset_subset_add_iff1: "j \ (i::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (repeat_mset (i-j) u + m \# n)" unfolding subset_mset_def repeat_mset_iterate_add by (simp add: iterate_add_eq_add_iff1 mset_subseteq_add_iff1[unfolded repeat_mset_iterate_add]) lemma mset_subset_add_iff2: "i \ (j::nat) \ (repeat_mset i u + m \# repeat_mset j u + n) = (m \# repeat_mset (j-i) u + n)" unfolding subset_mset_def repeat_mset_iterate_add by (simp add: iterate_add_eq_add_iff2 mset_subseteq_add_iff2[unfolded repeat_mset_iterate_add]) ML_file \multiset_simprocs.ML\ lemma add_mset_replicate_mset_safe[cancelation_simproc_pre]: \NO_MATCH {#} M \ add_mset a M = {#a#} + M\ by simp declare repeat_mset_iterate_add[cancelation_simproc_pre] declare iterate_add_distrib[cancelation_simproc_pre] declare repeat_mset_iterate_add[symmetric, cancelation_simproc_post] declare add_mset_not_empty[cancelation_simproc_eq_elim] empty_not_add_mset[cancelation_simproc_eq_elim] subset_mset.le_zero_eq[cancelation_simproc_eq_elim] empty_not_add_mset[cancelation_simproc_eq_elim] add_mset_not_empty[cancelation_simproc_eq_elim] subset_mset.le_zero_eq[cancelation_simproc_eq_elim] le_zero_eq[cancelation_simproc_eq_elim] simproc_setup mseteq_cancel ("(l::'a multiset) + m = n" | "(l::'a multiset) = m + n" | "add_mset a m = n" | "m = add_mset a n" | "replicate_mset p a = n" | "m = replicate_mset p a" | "repeat_mset p m = n" | "m = repeat_mset p m") = \K Cancel_Simprocs.eq_cancel\ simproc_setup msetsubset_cancel ("(l::'a multiset) + m \# n" | "(l::'a multiset) \# m + n" | "add_mset a m \# n" | "m \# add_mset a n" | "replicate_mset p r \# n" | "m \# replicate_mset p r" | "repeat_mset p m \# n" | "m \# repeat_mset p m") = \K Multiset_Simprocs.subset_cancel_msets\ simproc_setup msetsubset_eq_cancel ("(l::'a multiset) + m \# n" | "(l::'a multiset) \# m + n" | "add_mset a m \# n" | "m \# add_mset a n" | "replicate_mset p r \# n" | "m \# replicate_mset p r" | "repeat_mset p m \# n" | "m \# repeat_mset p m") = \K Multiset_Simprocs.subseteq_cancel_msets\ simproc_setup msetdiff_cancel ("((l::'a multiset) + m) - n" | "(l::'a multiset) - (m + n)" | "add_mset a m - n" | "m - add_mset a n" | "replicate_mset p r - n" | "m - replicate_mset p r" | "repeat_mset p m - n" | "m - repeat_mset p m") = \K Cancel_Simprocs.diff_cancel\ subsubsection \Conditionally complete lattice\ instantiation multiset :: (type) Inf begin lift_definition Inf_multiset :: "'a multiset set \ 'a multiset" is "\A i. if A = {} then 0 else Inf ((\f. f i) ` A)" proof - fix A :: "('a \ nat) set" assume *: "\f. f \ A \ finite {x. 0 < f x}" show \finite {i. 0 < (if A = {} then 0 else INF f\A. f i)}\ proof (cases "A = {}") case False then obtain f where "f \ A" by blast hence "{i. Inf ((\f. f i) ` A) > 0} \ {i. f i > 0}" by (auto intro: less_le_trans[OF _ cInf_lower]) moreover from \f \ A\ * have "finite \" by simp ultimately have "finite {i. Inf ((\f. f i) ` A) > 0}" by (rule finite_subset) with False show ?thesis by simp qed simp_all qed instance .. end lemma Inf_multiset_empty: "Inf {} = {#}" by transfer simp_all lemma count_Inf_multiset_nonempty: "A \ {} \ count (Inf A) x = Inf ((\X. count X x) ` A)" by transfer simp_all instantiation multiset :: (type) Sup begin definition Sup_multiset :: "'a multiset set \ 'a multiset" where "Sup_multiset A = (if A \ {} \ subset_mset.bdd_above A then Abs_multiset (\i. Sup ((\X. count X i) ` A)) else {#})" lemma Sup_multiset_empty: "Sup {} = {#}" by (simp add: Sup_multiset_def) lemma Sup_multiset_unbounded: "\ subset_mset.bdd_above A \ Sup A = {#}" by (simp add: Sup_multiset_def) instance .. end lemma bdd_above_multiset_imp_bdd_above_count: assumes "subset_mset.bdd_above (A :: 'a multiset set)" shows "bdd_above ((\X. count X x) ` A)" proof - from assms obtain Y where Y: "\X\A. X \# Y" by (meson subset_mset.bdd_above.E) hence "count X x \ count Y x" if "X \ A" for X using that by (auto intro: mset_subset_eq_count) thus ?thesis by (intro bdd_aboveI[of _ "count Y x"]) auto qed lemma bdd_above_multiset_imp_finite_support: assumes "A \ {}" "subset_mset.bdd_above (A :: 'a multiset set)" shows "finite (\X\A. {x. count X x > 0})" proof - from assms obtain Y where Y: "\X\A. X \# Y" by (meson subset_mset.bdd_above.E) hence "count X x \ count Y x" if "X \ A" for X x using that by (auto intro: mset_subset_eq_count) hence "(\X\A. {x. count X x > 0}) \ {x. count Y x > 0}" by safe (erule less_le_trans) moreover have "finite \" by simp ultimately show ?thesis by (rule finite_subset) qed lemma Sup_multiset_in_multiset: \finite {i. 0 < (SUP M\A. count M i)}\ if \A \ {}\ \subset_mset.bdd_above A\ proof - have "{i. Sup ((\X. count X i) ` A) > 0} \ (\X\A. {i. 0 < count X i})" proof safe fix i assume pos: "(SUP X\A. count X i) > 0" show "i \ (\X\A. {i. 0 < count X i})" proof (rule ccontr) assume "i \ (\X\A. {i. 0 < count X i})" hence "\X\A. count X i \ 0" by (auto simp: count_eq_zero_iff) with that have "(SUP X\A. count X i) \ 0" by (intro cSup_least bdd_above_multiset_imp_bdd_above_count) auto with pos show False by simp qed qed moreover from that have "finite \" by (rule bdd_above_multiset_imp_finite_support) ultimately show "finite {i. Sup ((\X. count X i) ` A) > 0}" by (rule finite_subset) qed lemma count_Sup_multiset_nonempty: \count (Sup A) x = (SUP X\A. count X x)\ if \A \ {}\ \subset_mset.bdd_above A\ using that by (simp add: Sup_multiset_def Sup_multiset_in_multiset count_Abs_multiset) interpretation subset_mset: conditionally_complete_lattice Inf Sup "(\#)" "(\#)" "(\#)" "(\#)" proof fix X :: "'a multiset" and A assume "X \ A" show "Inf A \# X" - proof (rule mset_subset_eqI) - fix x - from \X \ A\ have "A \ {}" by auto - hence "count (Inf A) x = (INF X\A. count X x)" - by (simp add: count_Inf_multiset_nonempty) - also from \X \ A\ have "\ \ count X x" - by (intro cInf_lower) simp_all - finally show "count (Inf A) x \ count X x" . - qed + by (metis \X \ A\ count_Inf_multiset_nonempty empty_iff image_eqI mset_subset_eqI wellorder_Inf_le1) next fix X :: "'a multiset" and A assume nonempty: "A \ {}" and le: "\Y. Y \ A \ X \# Y" show "X \# Inf A" proof (rule mset_subset_eqI) fix x from nonempty have "count X x \ (INF X\A. count X x)" by (intro cInf_greatest) (auto intro: mset_subset_eq_count le) also from nonempty have "\ = count (Inf A) x" by (simp add: count_Inf_multiset_nonempty) finally show "count X x \ count (Inf A) x" . qed next fix X :: "'a multiset" and A assume X: "X \ A" and bdd: "subset_mset.bdd_above A" show "X \# Sup A" proof (rule mset_subset_eqI) fix x from X have "A \ {}" by auto have "count X x \ (SUP X\A. count X x)" by (intro cSUP_upper X bdd_above_multiset_imp_bdd_above_count bdd) also from count_Sup_multiset_nonempty[OF \A \ {}\ bdd] have "(SUP X\A. count X x) = count (Sup A) x" by simp finally show "count X x \ count (Sup A) x" . qed next fix X :: "'a multiset" and A assume nonempty: "A \ {}" and ge: "\Y. Y \ A \ Y \# X" from ge have bdd: "subset_mset.bdd_above A" by blast show "Sup A \# X" proof (rule mset_subset_eqI) fix x from count_Sup_multiset_nonempty[OF \A \ {}\ bdd] have "count (Sup A) x = (SUP X\A. count X x)" . also from nonempty have "\ \ count X x" by (intro cSup_least) (auto intro: mset_subset_eq_count ge) finally show "count (Sup A) x \ count X x" . qed qed \ \FIXME: avoid junk stemming from type class interpretation\ lemma set_mset_Inf: assumes "A \ {}" shows "set_mset (Inf A) = (\X\A. set_mset X)" proof safe fix x X assume "x \# Inf A" "X \ A" hence nonempty: "A \ {}" by (auto simp: Inf_multiset_empty) from \x \# Inf A\ have "{#x#} \# Inf A" by auto also from \X \ A\ have "\ \# X" by (rule subset_mset.cInf_lower) simp_all finally show "x \# X" by simp next fix x assume x: "x \ (\X\A. set_mset X)" hence "{#x#} \# X" if "X \ A" for X using that by auto from assms and this have "{#x#} \# Inf A" by (rule subset_mset.cInf_greatest) thus "x \# Inf A" by simp qed lemma in_Inf_multiset_iff: assumes "A \ {}" shows "x \# Inf A \ (\X\A. x \# X)" proof - from assms have "set_mset (Inf A) = (\X\A. set_mset X)" by (rule set_mset_Inf) also have "x \ \ \ (\X\A. x \# X)" by simp finally show ?thesis . qed lemma in_Inf_multisetD: "x \# Inf A \ X \ A \ x \# X" by (subst (asm) in_Inf_multiset_iff) auto lemma set_mset_Sup: assumes "subset_mset.bdd_above A" shows "set_mset (Sup A) = (\X\A. set_mset X)" proof safe fix x assume "x \# Sup A" hence nonempty: "A \ {}" by (auto simp: Sup_multiset_empty) show "x \ (\X\A. set_mset X)" proof (rule ccontr) assume x: "x \ (\X\A. set_mset X)" have "count X x \ count (Sup A) x" if "X \ A" for X x using that by (intro mset_subset_eq_count subset_mset.cSup_upper assms) with x have "X \# Sup A - {#x#}" if "X \ A" for X using that by (auto simp: subseteq_mset_def algebra_simps not_in_iff) hence "Sup A \# Sup A - {#x#}" by (intro subset_mset.cSup_least nonempty) with \x \# Sup A\ show False - by (auto simp: subseteq_mset_def simp flip: count_greater_zero_iff - dest!: spec[of _ x]) + using mset_subset_diff_self by fastforce qed next fix x X assume "x \ set_mset X" "X \ A" hence "{#x#} \# X" by auto also have "X \# Sup A" by (intro subset_mset.cSup_upper \X \ A\ assms) finally show "x \ set_mset (Sup A)" by simp qed lemma in_Sup_multiset_iff: assumes "subset_mset.bdd_above A" shows "x \# Sup A \ (\X\A. x \# X)" -proof - - from assms have "set_mset (Sup A) = (\X\A. set_mset X)" by (rule set_mset_Sup) - also have "x \ \ \ (\X\A. x \# X)" by simp - finally show ?thesis . -qed + by (simp add: assms set_mset_Sup) lemma in_Sup_multisetD: assumes "x \# Sup A" shows "\X\A. x \# X" -proof - - have "subset_mset.bdd_above A" - by (rule ccontr) (insert assms, simp_all add: Sup_multiset_unbounded) - with assms show ?thesis by (simp add: in_Sup_multiset_iff) -qed + using Sup_multiset_unbounded assms in_Sup_multiset_iff by fastforce interpretation subset_mset: distrib_lattice "(\#)" "(\#)" "(\#)" "(\#)" proof fix A B C :: "'a multiset" show "A \# (B \# C) = A \# B \# (A \# C)" by (intro multiset_eqI) simp_all qed \ \FIXME: avoid junk stemming from type class interpretation\ subsubsection \Filter (with comprehension syntax)\ text \Multiset comprehension\ lift_definition filter_mset :: "('a \ bool) \ 'a multiset \ 'a multiset" is "\P M. \x. if P x then M x else 0" by (rule filter_preserves_multiset) syntax (ASCII) "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ :# _./ _#})") syntax "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{#_ \# _./ _#})") translations "{#x \# M. P#}" == "CONST filter_mset (\x. P) M" lemma count_filter_mset [simp]: "count (filter_mset P M) a = (if P a then count M a else 0)" by (simp add: filter_mset.rep_eq) lemma set_mset_filter [simp]: "set_mset (filter_mset P M) = {a \ set_mset M. P a}" by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_filter_mset) simp lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}" by (rule multiset_eqI) simp lemma filter_single_mset: "filter_mset P {#x#} = (if P x then {#x#} else {#})" by (rule multiset_eqI) simp lemma filter_union_mset [simp]: "filter_mset P (M + N) = filter_mset P M + filter_mset P N" by (rule multiset_eqI) simp lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N" by (rule multiset_eqI) simp lemma filter_inter_mset [simp]: "filter_mset P (M \# N) = filter_mset P M \# filter_mset P N" by (rule multiset_eqI) simp lemma filter_sup_mset[simp]: "filter_mset P (A \# B) = filter_mset P A \# filter_mset P B" by (rule multiset_eqI) simp lemma filter_mset_add_mset [simp]: "filter_mset P (add_mset x A) = (if P x then add_mset x (filter_mset P A) else filter_mset P A)" by (auto simp: multiset_eq_iff) lemma multiset_filter_subset[simp]: "filter_mset f M \# M" by (simp add: mset_subset_eqI) lemma multiset_filter_mono: assumes "A \# B" shows "filter_mset f A \# filter_mset f B" -proof - - from assms[unfolded mset_subset_eq_exists_conv] - obtain C where B: "B = A + C" by auto - show ?thesis unfolding B by auto -qed + by (metis assms filter_sup_mset subset_mset.order_iff) lemma filter_mset_eq_conv: "filter_mset P M = N \ N \# M \ (\b\#N. P b) \ (\a\#M - N. \ P a)" (is "?P \ ?Q") proof assume ?P then show ?Q by auto (simp add: multiset_eq_iff in_diff_count) next assume ?Q then obtain Q where M: "M = N + Q" by (auto simp add: mset_subset_eq_exists_conv) then have MN: "M - N = Q" by simp show ?P proof (rule multiset_eqI) fix a from \?Q\ MN have *: "\ P a \ a \# N" "P a \ a \# Q" by auto show "count (filter_mset P M) a = count N a" proof (cases "a \# M") case True with * show ?thesis by (simp add: not_in_iff M) next case False then have "count M a = 0" by (simp add: not_in_iff) with M show ?thesis by simp qed qed qed lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x \# M. Q x \ P x#}" by (auto simp: multiset_eq_iff) lemma filter_mset_True[simp]: "{#y \# M. True#} = M" and filter_mset_False[simp]: "{#y \# M. False#} = {#}" by (auto simp: multiset_eq_iff) lemma filter_mset_cong0: assumes "\x. x \# M \ f x \ g x" shows "filter_mset f M = filter_mset g M" proof (rule subset_mset.antisym; unfold subseteq_mset_def; rule allI) fix x show "count (filter_mset f M) x \ count (filter_mset g M) x" using assms by (cases "x \# M") (simp_all add: not_in_iff) next fix x show "count (filter_mset g M) x \ count (filter_mset f M) x" using assms by (cases "x \# M") (simp_all add: not_in_iff) qed lemma filter_mset_cong: assumes "M = M'" and "\x. x \# M' \ f x \ g x" shows "filter_mset f M = filter_mset g M'" unfolding \M = M'\ using assms by (auto intro: filter_mset_cong0) lemma filter_eq_replicate_mset: "{#y \# D. y = x#} = replicate_mset (count D x) x" by (induct D) (simp add: multiset_eqI) subsubsection \Size\ definition wcount where "wcount f M = (\x. count M x * Suc (f x))" lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" by (auto simp: wcount_def add_mult_distrib) lemma wcount_add_mset: "wcount f (add_mset x M) a = (if x = a then Suc (f a) else 0) + wcount f M a" unfolding add_mset_add_single[of _ M] wcount_union by (auto simp: wcount_def) definition size_multiset :: "('a \ nat) \ 'a multiset \ nat" where "size_multiset f M = sum (wcount f M) (set_mset M)" lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] instantiation multiset :: (type) size begin definition size_multiset where size_multiset_overloaded_def: "size_multiset = Multiset.size_multiset (\_. 0)" instance .. end lemmas size_multiset_overloaded_eq = size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] lemma size_multiset_empty [simp]: "size_multiset f {#} = 0" -by (simp add: size_multiset_def) + by (simp add: size_multiset_def) lemma size_empty [simp]: "size {#} = 0" -by (simp add: size_multiset_overloaded_def) + by (simp add: size_multiset_overloaded_def) lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)" -by (simp add: size_multiset_eq) + by (simp add: size_multiset_eq) lemma size_single: "size {#b#} = 1" -by (simp add: size_multiset_overloaded_def size_multiset_single) + by (simp add: size_multiset_overloaded_def size_multiset_single) lemma sum_wcount_Int: "finite A \ sum (wcount f N) (A \ set_mset N) = sum (wcount f N) A" by (induct rule: finite_induct) (simp_all add: Int_insert_left wcount_def count_eq_zero_iff) lemma size_multiset_union [simp]: "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" -apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union) -apply (subst Int_commute) -apply (simp add: sum_wcount_Int) -done + apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union) + by (metis add_implies_diff finite_set_mset inf.commute sum_wcount_Int) lemma size_multiset_add_mset [simp]: "size_multiset f (add_mset a M) = Suc (f a) + size_multiset f M" - unfolding add_mset_add_single[of _ M] size_multiset_union by (auto simp: size_multiset_single) + by (metis add.commute add_mset_add_single size_multiset_single size_multiset_union) lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)" -by (simp add: size_multiset_overloaded_def wcount_add_mset) + by (simp add: size_multiset_overloaded_def wcount_add_mset) lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" -by (auto simp add: size_multiset_overloaded_def) + by (auto simp add: size_multiset_overloaded_def) lemma size_multiset_eq_0_iff_empty [iff]: "size_multiset f M = 0 \ M = {#}" by (auto simp add: size_multiset_eq count_eq_zero_iff) lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" by (auto simp add: size_multiset_overloaded_def) lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" -by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) + by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) lemma size_eq_Suc_imp_elem: "size M = Suc n \ \a. a \# M" -apply (unfold size_multiset_overloaded_eq) -apply (drule sum_SucD) -apply auto -done + using all_not_in_conv by fastforce lemma size_eq_Suc_imp_eq_union: assumes "size M = Suc n" shows "\a N. M = add_mset a N" -proof - - from assms obtain a where "a \# M" - by (erule size_eq_Suc_imp_elem [THEN exE]) - then have "M = add_mset a (M - {#a#})" by simp - then show ?thesis by blast -qed + by (metis assms insert_DiffM size_eq_Suc_imp_elem) lemma size_mset_mono: fixes A B :: "'a multiset" assumes "A \# B" shows "size A \ size B" proof - from assms[unfolded mset_subset_eq_exists_conv] obtain C where B: "B = A + C" by auto show ?thesis unfolding B by (induct C) auto qed lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \ size M" -by (rule size_mset_mono[OF multiset_filter_subset]) + by (rule size_mset_mono[OF multiset_filter_subset]) lemma size_Diff_submset: "M \# M' \ size (M' - M) = size M' - size(M::'a multiset)" by (metis add_diff_cancel_left' size_union mset_subset_eq_exists_conv) lemma size_lt_imp_ex_count_lt: "size M < size N \ \x \# N. count M x < count N x" by (metis count_eq_zero_iff leD not_le_imp_less not_less_zero size_mset_mono subseteq_mset_def) subsection \Induction and case splits\ theorem multiset_induct [case_names empty add, induct type: multiset]: assumes empty: "P {#}" assumes add: "\x M. P M \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case 0 thus "P M" by (simp add: empty) next case (Suc k) obtain N x where "M = add_mset x N" using \Suc k = size M\ [symmetric] using size_eq_Suc_imp_eq_union by fast with Suc add show "P M" by simp qed lemma multiset_induct_min[case_names empty add]: fixes M :: "'a::linorder multiset" assumes empty: "P {#}" and add: "\x M. P M \ (\y \# M. y \ x) \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case (Suc k) note ih = this(1) and Sk_eq_sz_M = this(2) let ?y = "Min_mset M" let ?N = "M - {#?y#}" have M: "M = add_mset ?y ?N" by (metis Min_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero set_mset_eq_empty_iff size_empty) show ?case by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, meson Min_le finite_set_mset in_diffD) qed (simp add: empty) lemma multiset_induct_max[case_names empty add]: fixes M :: "'a::linorder multiset" assumes empty: "P {#}" and add: "\x M. P M \ (\y \# M. y \ x) \ P (add_mset x M)" shows "P M" proof (induct "size M" arbitrary: M) case (Suc k) note ih = this(1) and Sk_eq_sz_M = this(2) let ?y = "Max_mset M" let ?N = "M - {#?y#}" have M: "M = add_mset ?y ?N" by (metis Max_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero set_mset_eq_empty_iff size_empty) show ?case by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, meson Max_ge finite_set_mset in_diffD) qed (simp add: empty) lemma multi_nonempty_split: "M \ {#} \ \A a. M = add_mset a A" -by (induct M) auto + by (induct M) auto lemma multiset_cases [cases type]: - obtains (empty) "M = {#}" - | (add) x N where "M = add_mset x N" + obtains (empty) "M = {#}" | (add) x N where "M = add_mset x N" by (induct M) simp_all lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" -by (cases "B = {#}") (auto dest: multi_member_split) + by (cases "B = {#}") (auto dest: multi_member_split) lemma union_filter_mset_complement[simp]: "\x. P x = (\ Q x) \ filter_mset P M + filter_mset Q M = M" -by (subst multiset_eq_iff) auto + by (subst multiset_eq_iff) auto lemma multiset_partition: "M = {#x \# M. P x#} + {#x \# M. \ P x#}" -by simp + by simp lemma mset_subset_size: "A \# B \ size A < size B" proof (induct A arbitrary: B) case empty then show ?case using nonempty_has_size by auto next case (add x A) have "add_mset x A \# B" by (meson add.prems subset_mset_def) then show ?case - by (metis (no_types) add.prems add.right_neutral add_diff_cancel_left' leD nat_neq_iff - size_Diff_submset size_eq_0_iff_empty size_mset_mono subset_mset.le_iff_add subset_mset_def) + using add.prems subset_mset.less_eqE by fastforce qed lemma size_1_singleton_mset: "size M = 1 \ \a. M = {#a#}" by (cases M) auto lemma set_mset_subset_singletonD: assumes "set_mset A \ {x}" shows "A = replicate_mset (size A) x" using assms by (induction A) auto subsubsection \Strong induction and subset induction for multisets\ text \Well-foundedness of strict subset relation\ lemma wf_subset_mset_rel: "wf {(M, N :: 'a multiset). M \# N}" -apply (rule wf_measure [THEN wf_subset, where f1=size]) -apply (clarsimp simp: measure_def inv_image_def mset_subset_size) -done + using mset_subset_size wfP_def wfP_if_convertible_to_nat by blast lemma wfP_subset_mset[simp]: "wfP (\#)" by (rule wf_subset_mset_rel[to_pred]) lemma full_multiset_induct [case_names less]: -assumes ih: "\B. \(A::'a multiset). A \# B \ P A \ P B" -shows "P B" -apply (rule wf_subset_mset_rel [THEN wf_induct]) -apply (rule ih, auto) -done + assumes ih: "\B. \(A::'a multiset). A \# B \ P A \ P B" + shows "P B" + apply (rule wf_subset_mset_rel [THEN wf_induct]) + apply (rule ih, auto) + done lemma multi_subset_induct [consumes 2, case_names empty add]: assumes "F \# A" and empty: "P {#}" and insert: "\a F. a \# A \ P F \ P (add_mset a F)" shows "P F" proof - from \F \# A\ show ?thesis proof (induct F) show "P {#}" by fact next fix x F assume P: "F \# A \ P F" and i: "add_mset x F \# A" show "P (add_mset x F)" proof (rule insert) from i show "x \# A" by (auto dest: mset_subset_eq_insertD) from i have "F \# A" by (auto dest: mset_subset_eq_insertD) with P show "P F" . qed qed qed subsection \Least and greatest elements\ context begin qualified lemma assumes "M \ {#}" and "transp_on (set_mset M) R" and "totalp_on (set_mset M) R" shows bex_least_element: "(\l \# M. \x \# M. x \ l \ R l x)" and bex_greatest_element: "(\g \# M. \x \# M. x \ g \ R x g)" using assms by (auto intro: Finite_Set.bex_least_element Finite_Set.bex_greatest_element) end subsection \The fold combinator\ definition fold_mset :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b" where "fold_mset f s M = Finite_Set.fold (\x. f x ^^ count M x) s (set_mset M)" lemma fold_mset_empty [simp]: "fold_mset f s {#} = s" by (simp add: fold_mset_def) lemma fold_mset_single [simp]: "fold_mset f s {#x#} = f x s" by (simp add: fold_mset_def) context comp_fun_commute begin lemma fold_mset_add_mset [simp]: "fold_mset f s (add_mset x M) = f x (fold_mset f s M)" proof - interpret mset: comp_fun_commute "\y. f y ^^ count M y" by (fact comp_fun_commute_funpow) interpret mset_union: comp_fun_commute "\y. f y ^^ count (add_mset x M) y" by (fact comp_fun_commute_funpow) show ?thesis proof (cases "x \ set_mset M") case False then have *: "count (add_mset x M) x = 1" by (simp add: not_in_iff) from False have "Finite_Set.fold (\y. f y ^^ count (add_mset x M) y) s (set_mset M) = Finite_Set.fold (\y. f y ^^ count M y) s (set_mset M)" by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow) with False * show ?thesis by (simp add: fold_mset_def del: count_add_mset) next case True define N where "N = set_mset M - {x}" from N_def True have *: "set_mset M = insert x N" "x \ N" "finite N" by auto then have "Finite_Set.fold (\y. f y ^^ count (add_mset x M) y) s N = Finite_Set.fold (\y. f y ^^ count M y) s N" by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow) with * show ?thesis by (simp add: fold_mset_def del: count_add_mset) simp qed qed lemma fold_mset_fun_left_comm: "f x (fold_mset f s M) = fold_mset f (f x s) M" by (induct M) (simp_all add: fun_left_comm) lemma fold_mset_union [simp]: "fold_mset f s (M + N) = fold_mset f (fold_mset f s M) N" by (induct M) (simp_all add: fold_mset_fun_left_comm) lemma fold_mset_fusion: assumes "comp_fun_commute g" and *: "\x y. h (g x y) = f x (h y)" shows "h (fold_mset g w A) = fold_mset f (h w) A" proof - interpret comp_fun_commute g by (fact assms) from * show ?thesis by (induct A) auto qed end lemma union_fold_mset_add_mset: "A + B = fold_mset add_mset A B" proof - interpret comp_fun_commute add_mset by standard auto show ?thesis by (induction B) auto qed text \ A note on code generation: When defining some function containing a subterm \<^term>\fold_mset F\, code generation is not automatic. When interpreting locale \left_commutative\ with \F\, the would be code thms for \<^const>\fold_mset\ become thms like \<^term>\fold_mset F z {#} = z\ where \F\ is not a pattern but contains defined symbols, i.e.\ is not a code thm. Hence a separate constant with its own code thms needs to be introduced for \F\. See the image operator below. \ subsection \Image\ definition image_mset :: "('a \ 'b) \ 'a multiset \ 'b multiset" where "image_mset f = fold_mset (add_mset \ f) {#}" lemma comp_fun_commute_mset_image: "comp_fun_commute (add_mset \ f)" by unfold_locales (simp add: fun_eq_iff) lemma image_mset_empty [simp]: "image_mset f {#} = {#}" by (simp add: image_mset_def) lemma image_mset_single: "image_mset f {#x#} = {#f x#}" by (simp add: comp_fun_commute.fold_mset_add_mset comp_fun_commute_mset_image image_mset_def) lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" proof - interpret comp_fun_commute "add_mset \ f" by (fact comp_fun_commute_mset_image) show ?thesis by (induct N) (simp_all add: image_mset_def) qed corollary image_mset_add_mset [simp]: "image_mset f (add_mset a M) = add_mset (f a) (image_mset f M)" unfolding image_mset_union add_mset_add_single[of a M] by (simp add: image_mset_single) lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)" by (induct M) simp_all lemma size_image_mset [simp]: "size (image_mset f M) = size M" by (induct M) simp_all lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" by (cases M) auto lemma image_mset_If: "image_mset (\x. if P x then f x else g x) A = image_mset f (filter_mset P A) + image_mset g (filter_mset (\x. \P x) A)" by (induction A) auto lemma image_mset_Diff: assumes "B \# A" shows "image_mset f (A - B) = image_mset f A - image_mset f B" proof - have "image_mset f (A - B + B) = image_mset f (A - B) + image_mset f B" by simp also from assms have "A - B + B = A" by (simp add: subset_mset.diff_add) finally show ?thesis by simp qed lemma count_image_mset: \count (image_mset f A) x = (\y\f -` {x} \ set_mset A. count A y)\ proof (induction A) case empty then show ?case by simp next case (add x A) moreover have *: "(if x = y then Suc n else n) = n + (if x = y then 1 else 0)" for n y by simp ultimately show ?case by (auto simp: sum.distrib intro!: sum.mono_neutral_left) qed lemma count_image_mset': \count (image_mset f X) y = (\x | x \# X \ y = f x. count X x)\ by (auto simp add: count_image_mset simp flip: singleton_conv2 simp add: Collect_conj_eq ac_simps) lemma image_mset_subseteq_mono: "A \# B \ image_mset f A \# image_mset f B" by (metis image_mset_union subset_mset.le_iff_add) lemma image_mset_subset_mono: "M \# N \ image_mset f M \# image_mset f N" by (metis (no_types) Diff_eq_empty_iff_mset image_mset_Diff image_mset_is_empty_iff image_mset_subseteq_mono subset_mset.less_le_not_le) syntax (ASCII) "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ :# _#})") syntax "_comprehension_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" ("({#_/. _ \# _#})") translations "{#e. x \# M#}" \ "CONST image_mset (\x. e) M" syntax (ASCII) "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ :# _./ _#})") syntax "_comprehension_mset'" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" ("({#_/ | _ \# _./ _#})") translations "{#e | x\#M. P#}" \ "{#e. x \# {# x\#M. P#}#}" text \ This allows to write not just filters like \<^term>\{#x\#M. x but also images like \<^term>\{#x+x. x\#M #}\ and @{term [source] "{#x+x|x\#M. x\{#x+x|x\#M. x. \ lemma in_image_mset: "y \# {#f x. x \# M#} \ y \ f ` set_mset M" by simp functor image_mset: image_mset proof - fix f g show "image_mset f \ image_mset g = image_mset (f \ g)" proof fix A show "(image_mset f \ image_mset g) A = image_mset (f \ g) A" by (induct A) simp_all qed show "image_mset id = id" proof fix A show "image_mset id A = id A" by (induct A) simp_all qed qed declare image_mset.id [simp] image_mset.identity [simp] lemma image_mset_id[simp]: "image_mset id x = x" unfolding id_def by auto lemma image_mset_cong: "(\x. x \# M \ f x = g x) \ {#f x. x \# M#} = {#g x. x \# M#}" by (induct M) auto lemma image_mset_cong_pair: "(\x y. (x, y) \# M \ f x y = g x y) \ {#f x y. (x, y) \# M#} = {#g x y. (x, y) \# M#}" by (metis image_mset_cong split_cong) lemma image_mset_const_eq: "{#c. a \# M#} = replicate_mset (size M) c" by (induct M) simp_all lemma image_mset_filter_mset_swap: "image_mset f (filter_mset (\x. P (f x)) M) = filter_mset P (image_mset f M)" by (induction M rule: multiset_induct) simp_all lemma image_mset_eq_plusD: "image_mset f A = B + C \ \B' C'. A = B' + C' \ B = image_mset f B' \ C = image_mset f C'" proof (induction A arbitrary: B C) case empty thus ?case by simp next case (add x A) show ?case proof (cases "f x \# B") case True with add.prems have "image_mset f A = (B - {#f x#}) + C" by (metis add_mset_remove_trivial image_mset_add_mset mset_subset_eq_single subset_mset.add_diff_assoc2) thus ?thesis using add.IH add.prems by force next case False with add.prems have "image_mset f A = B + (C - {#f x#})" by (metis diff_single_eq_union diff_union_single_conv image_mset_add_mset union_iff union_single_eq_member) then show ?thesis using add.IH add.prems by force qed qed lemma image_mset_eq_image_mset_plusD: assumes "image_mset f A = image_mset f B + C" and inj_f: "inj_on f (set_mset A \ set_mset B)" shows "\C'. A = B + C' \ C = image_mset f C'" using assms proof (induction A arbitrary: B C) case empty thus ?case by simp next case (add x A) show ?case proof (cases "x \# B") case True with add.prems have "image_mset f A = image_mset f (B - {#x#}) + C" - by (smt (verit, del_insts) add.left_commute add_cancel_right_left diff_union_cancelL - diff_union_single_conv image_mset_union union_mset_add_mset_left - union_mset_add_mset_right) + by (smt (verit) add_mset_add_mset_same_iff image_mset_add_mset insert_DiffM union_mset_add_mset_left) with add.IH have "\M3'. A = B - {#x#} + M3' \ image_mset f M3' = C" by (smt (verit, del_insts) True Un_insert_left Un_insert_right add.prems(2) inj_on_insert insert_DiffM set_mset_add_mset_insert) with True show ?thesis by auto next case False with add.prems(2) have "f x \# image_mset f B" by auto with add.prems(1) have "image_mset f A = image_mset f B + (C - {#f x#})" by (metis (no_types, lifting) diff_union_single_conv image_eqI image_mset_Diff image_mset_single mset_subset_eq_single set_image_mset union_iff union_single_eq_diff union_single_eq_member) with add.prems(2) add.IH have "\M3'. A = B + M3' \ C - {#f x#} = image_mset f M3'" by auto then show ?thesis by (metis add.prems(1) add_diff_cancel_left' image_mset_Diff mset_subset_eq_add_left union_mset_add_mset_right) qed qed lemma image_mset_eq_plus_image_msetD: "image_mset f A = B + image_mset f C \ inj_on f (set_mset A \ set_mset C) \ \B'. A = B' + C \ B = image_mset f B'" unfolding add.commute[of B] add.commute[of _ C] by (rule image_mset_eq_image_mset_plusD; assumption) subsection \Further conversions\ primrec mset :: "'a list \ 'a multiset" where "mset [] = {#}" | "mset (a # x) = add_mset a (mset x)" lemma in_multiset_in_set: "x \# mset xs \ x \ set xs" by (induct xs) simp_all lemma count_mset: "count (mset xs) x = length (filter (\y. x = y) xs)" by (induct xs) simp_all lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])" by (induct x) auto lemma mset_zero_iff_right[simp]: "({#} = mset x) = (x = [])" by (induct x) auto lemma count_mset_gt_0: "x \ set xs \ count (mset xs) x > 0" by (induction xs) auto lemma count_mset_0_iff [simp]: "count (mset xs) x = 0 \ x \ set xs" by (induction xs) auto lemma mset_single_iff[iff]: "mset xs = {#x#} \ xs = [x]" by (cases xs) auto lemma mset_single_iff_right[iff]: "{#x#} = mset xs \ xs = [x]" by (cases xs) auto lemma set_mset_mset[simp]: "set_mset (mset xs) = set xs" by (induct xs) auto lemma set_mset_comp_mset [simp]: "set_mset \ mset = set" by (simp add: fun_eq_iff) lemma size_mset [simp]: "size (mset xs) = length xs" by (induct xs) simp_all lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys" by (induct xs arbitrary: ys) auto lemma mset_filter[simp]: "mset (filter P xs) = {#x \# mset xs. P x #}" by (induct xs) simp_all lemma mset_rev [simp]: "mset (rev xs) = mset xs" by (induct xs) simp_all lemma surj_mset: "surj mset" unfolding surj_def proof (rule allI) fix M show "\xs. M = mset xs" by (induction M) (auto intro: exI[of _ "_ # _"]) qed lemma distinct_count_atmost_1: "distinct x = (\a. count (mset x) a = (if a \ set x then 1 else 0))" proof (induct x) case Nil then show ?case by simp next case (Cons x xs) show ?case (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs using Cons by simp next assume ?rhs then have "x \ set xs" by (simp split: if_splits) moreover from \?rhs\ have "(\a. count (mset xs) a = (if a \ set xs then 1 else 0))" by (auto split: if_splits simp add: count_eq_zero_iff) ultimately show ?lhs using Cons by simp qed qed lemma mset_eq_setD: assumes "mset xs = mset ys" shows "set xs = set ys" proof - from assms have "set_mset (mset xs) = set_mset (mset ys)" by simp then show ?thesis by simp qed lemma set_eq_iff_mset_eq_distinct: \distinct x \ distinct y \ set x = set y \ mset x = mset y\ by (auto simp: multiset_eq_iff distinct_count_atmost_1) lemma set_eq_iff_mset_remdups_eq: \set x = set y \ mset (remdups x) = mset (remdups y)\ -apply (rule iffI) -apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1]) -apply (drule distinct_remdups [THEN distinct_remdups - [THEN set_eq_iff_mset_eq_distinct [THEN iffD2]]]) -apply simp -done + using set_eq_iff_mset_eq_distinct by fastforce lemma mset_eq_imp_distinct_iff: \distinct xs \ distinct ys\ if \mset xs = mset ys\ using that by (auto simp add: distinct_count_atmost_1 dest: mset_eq_setD) lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" proof (induct ls arbitrary: i) case Nil then show ?case by simp next case Cons then show ?case by (cases i) auto qed lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}" by (induct xs) (auto simp add: multiset_eq_iff) lemma mset_eq_length: assumes "mset xs = mset ys" shows "length xs = length ys" using assms by (metis size_mset) lemma mset_eq_length_filter: assumes "mset xs = mset ys" shows "length (filter (\x. z = x) xs) = length (filter (\y. z = y) ys)" using assms by (metis count_mset) lemma fold_multiset_equiv: \List.fold f xs = List.fold f ys\ if f: \\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x\ and \mset xs = mset ys\ using f \mset xs = mset ys\ [symmetric] proof (induction xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) then have *: \set ys = set (x # xs)\ by (blast dest: mset_eq_setD) have \\x y. x \ set ys \ y \ set ys \ f x \ f y = f y \ f x\ by (rule Cons.prems(1)) (simp_all add: *) moreover from * have \x \ set ys\ by simp ultimately have \List.fold f ys = List.fold f (remove1 x ys) \ f x\ by (fact fold_remove1_split) moreover from Cons.prems have \List.fold f xs = List.fold f (remove1 x ys)\ by (auto intro: Cons.IH) ultimately show ?case by simp qed lemma fold_permuted_eq: \List.fold (\) xs z = List.fold (\) ys z\ if \mset xs = mset ys\ and \P z\ and P: \\x z. x \ set xs \ P z \ P (x \ z)\ and f: \\x y z. x \ set xs \ y \ set xs \ P z \ x \ (y \ z) = y \ (x \ z)\ for f (infixl \\\ 70) using \P z\ P f \mset xs = mset ys\ [symmetric] proof (induction xs arbitrary: ys z) case Nil then show ?case by simp next case (Cons x xs) then have *: \set ys = set (x # xs)\ by (blast dest: mset_eq_setD) have \P z\ by (fact Cons.prems(1)) moreover have \\x z. x \ set ys \ P z \ P (x \ z)\ by (rule Cons.prems(2)) (simp_all add: *) moreover have \\x y z. x \ set ys \ y \ set ys \ P z \ x \ (y \ z) = y \ (x \ z)\ by (rule Cons.prems(3)) (simp_all add: *) moreover from * have \x \ set ys\ by simp ultimately have \fold (\) ys z = fold (\) (remove1 x ys) (x \ z)\ by (induction ys arbitrary: z) auto moreover from Cons.prems have \fold (\) xs (x \ z) = fold (\) (remove1 x ys) (x \ z)\ by (auto intro: Cons.IH) ultimately show ?case by simp qed lemma mset_shuffles: "zs \ shuffles xs ys \ mset zs = mset xs + mset ys" by (induction xs ys arbitrary: zs rule: shuffles.induct) auto lemma mset_insort [simp]: "mset (insort x xs) = add_mset x (mset xs)" by (induct xs) simp_all lemma mset_map[simp]: "mset (map f xs) = image_mset f (mset xs)" by (induct xs) simp_all global_interpretation mset_set: folding add_mset "{#}" defines mset_set = "folding_on.F add_mset {#}" by standard (simp add: fun_eq_iff) lemma sum_multiset_singleton [simp]: "sum (\n. {#n#}) A = mset_set A" by (induction A rule: infinite_finite_induct) auto lemma count_mset_set [simp]: "finite A \ x \ A \ count (mset_set A) x = 1" (is "PROP ?P") "\ finite A \ count (mset_set A) x = 0" (is "PROP ?Q") "x \ A \ count (mset_set A) x = 0" (is "PROP ?R") proof - have *: "count (mset_set A) x = 0" if "x \ A" for A proof (cases "finite A") case False then show ?thesis by simp next case True from True \x \ A\ show ?thesis by (induct A) auto qed then show "PROP ?P" "PROP ?Q" "PROP ?R" by (auto elim!: Set.set_insert) qed \ \TODO: maybe define \<^const>\mset_set\ also in terms of \<^const>\Abs_multiset\\ lemma elem_mset_set[simp, intro]: "finite A \ x \# mset_set A \ x \ A" by (induct A rule: finite_induct) simp_all lemma mset_set_Union: "finite A \ finite B \ A \ B = {} \ mset_set (A \ B) = mset_set A + mset_set B" by (induction A rule: finite_induct) auto lemma filter_mset_mset_set [simp]: "finite A \ filter_mset P (mset_set A) = mset_set {x\A. P x}" proof (induction A rule: finite_induct) case (insert x A) from insert.hyps have "filter_mset P (mset_set (insert x A)) = filter_mset P (mset_set A) + mset_set (if P x then {x} else {})" by simp also have "filter_mset P (mset_set A) = mset_set {x\A. P x}" by (rule insert.IH) also from insert.hyps have "\ + mset_set (if P x then {x} else {}) = mset_set ({x \ A. P x} \ (if P x then {x} else {}))" (is "_ = mset_set ?A") by (intro mset_set_Union [symmetric]) simp_all also from insert.hyps have "?A = {y\insert x A. P y}" by auto finally show ?case . qed simp_all lemma mset_set_Diff: assumes "finite A" "B \ A" shows "mset_set (A - B) = mset_set A - mset_set B" proof - from assms have "mset_set ((A - B) \ B) = mset_set (A - B) + mset_set B" by (intro mset_set_Union) (auto dest: finite_subset) also from assms have "A - B \ B = A" by blast finally show ?thesis by simp qed lemma mset_set_set: "distinct xs \ mset_set (set xs) = mset xs" by (induction xs) simp_all lemma count_mset_set': "count (mset_set A) x = (if finite A \ x \ A then 1 else 0)" by auto lemma subset_imp_msubset_mset_set: assumes "A \ B" "finite B" shows "mset_set A \# mset_set B" proof (rule mset_subset_eqI) fix x :: 'a from assms have "finite A" by (rule finite_subset) with assms show "count (mset_set A) x \ count (mset_set B) x" by (cases "x \ A"; cases "x \ B") auto qed lemma mset_set_set_mset_msubset: "mset_set (set_mset A) \# A" proof (rule mset_subset_eqI) fix x show "count (mset_set (set_mset A)) x \ count A x" by (cases "x \# A") simp_all qed lemma mset_set_upto_eq_mset_upto: \mset_set {.. by (induction n) (auto simp: ac_simps lessThan_Suc) context linorder begin definition sorted_list_of_multiset :: "'a multiset \ 'a list" where "sorted_list_of_multiset M = fold_mset insort [] M" lemma sorted_list_of_multiset_empty [simp]: "sorted_list_of_multiset {#} = []" by (simp add: sorted_list_of_multiset_def) lemma sorted_list_of_multiset_singleton [simp]: "sorted_list_of_multiset {#x#} = [x]" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_multiset_def) qed lemma sorted_list_of_multiset_insert [simp]: "sorted_list_of_multiset (add_mset x M) = List.insort x (sorted_list_of_multiset M)" proof - interpret comp_fun_commute insort by (fact comp_fun_commute_insort) show ?thesis by (simp add: sorted_list_of_multiset_def) qed end lemma mset_sorted_list_of_multiset[simp]: "mset (sorted_list_of_multiset M) = M" by (induct M) simp_all lemma sorted_list_of_multiset_mset[simp]: "sorted_list_of_multiset (mset xs) = sort xs" by (induct xs) simp_all lemma finite_set_mset_mset_set[simp]: "finite A \ set_mset (mset_set A) = A" by auto lemma mset_set_empty_iff: "mset_set A = {#} \ A = {} \ infinite A" using finite_set_mset_mset_set by fastforce lemma infinite_set_mset_mset_set: "\ finite A \ set_mset (mset_set A) = {}" by simp lemma set_sorted_list_of_multiset [simp]: "set (sorted_list_of_multiset M) = set_mset M" -by (induct M) (simp_all add: set_insort_key) + by (induct M) (simp_all add: set_insort_key) lemma sorted_list_of_mset_set [simp]: "sorted_list_of_multiset (mset_set A) = sorted_list_of_set A" -by (cases "finite A") (induct A rule: finite_induct, simp_all) + by (cases "finite A") (induct A rule: finite_induct, simp_all) lemma mset_upt [simp]: "mset [m.. {#the (map_of xs i). i \# mset (map fst xs)#} = mset (map snd xs)" proof (induction xs) case (Cons x xs) have "{#the (map_of (x # xs) i). i \# mset (map fst (x # xs))#} = add_mset (snd x) {#the (if i = fst x then Some (snd x) else map_of xs i). i \# mset (map fst xs)#}" (is "_ = add_mset _ ?A") by simp also from Cons.prems have "?A = {#the (map_of xs i). i :# mset (map fst xs)#}" by (cases x, intro image_mset_cong) (auto simp: in_multiset_in_set) also from Cons.prems have "\ = mset (map snd xs)" by (intro Cons.IH) simp_all finally show ?case by simp qed simp_all lemma msubset_mset_set_iff[simp]: assumes "finite A" "finite B" shows "mset_set A \# mset_set B \ A \ B" using assms set_mset_mono subset_imp_msubset_mset_set by fastforce lemma mset_set_eq_iff[simp]: assumes "finite A" "finite B" shows "mset_set A = mset_set B \ A = B" using assms by (fastforce dest: finite_set_mset_mset_set) lemma image_mset_mset_set: \<^marker>\contributor \Lukas Bulwahn\\ assumes "inj_on f A" shows "image_mset f (mset_set A) = mset_set (f ` A)" proof cases assume "finite A" from this \inj_on f A\ show ?thesis by (induct A) auto next assume "infinite A" from this \inj_on f A\ have "infinite (f ` A)" using finite_imageD by blast from \infinite A\ \infinite (f ` A)\ show ?thesis by simp qed subsection \More properties of the replicate, repeat, and image operations\ lemma in_replicate_mset[simp]: "x \# replicate_mset n y \ n > 0 \ x = y" unfolding replicate_mset_def by (induct n) auto lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})" - by (auto split: if_splits) + by auto lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" by (induct n, simp_all) lemma size_repeat_mset [simp]: "size (repeat_mset n A) = n * size A" by (induction n) auto lemma size_multiset_sum [simp]: "size (\x\A. f x :: 'a multiset) = (\x\A. size (f x))" by (induction A rule: infinite_finite_induct) auto lemma size_multiset_sum_list [simp]: "size (\X\Xs. X :: 'a multiset) = (\X\Xs. size X)" by (induction Xs) auto lemma count_le_replicate_mset_subset_eq: "n \ count M x \ replicate_mset n x \# M" by (auto simp add: mset_subset_eqI) (metis count_replicate_mset subseteq_mset_def) lemma replicate_count_mset_eq_filter_eq: "replicate (count (mset xs) k) k = filter (HOL.eq k) xs" by (induct xs) auto lemma replicate_mset_eq_empty_iff [simp]: "replicate_mset n a = {#} \ n = 0" by (induct n) simp_all lemma replicate_mset_eq_iff: "replicate_mset m a = replicate_mset n b \ m = 0 \ n = 0 \ m = n \ a = b" by (auto simp add: multiset_eq_iff) lemma repeat_mset_cancel1: "repeat_mset a A = repeat_mset a B \ A = B \ a = 0" by (auto simp: multiset_eq_iff) lemma repeat_mset_cancel2: "repeat_mset a A = repeat_mset b A \ a = b \ A = {#}" by (auto simp: multiset_eq_iff) lemma repeat_mset_eq_empty_iff: "repeat_mset n A = {#} \ n = 0 \ A = {#}" by (cases n) auto lemma image_replicate_mset [simp]: "image_mset f (replicate_mset n a) = replicate_mset n (f a)" by (induct n) simp_all lemma replicate_mset_msubseteq_iff: "replicate_mset m a \# replicate_mset n b \ m = 0 \ a = b \ m \ n" by (cases m) (auto simp: insert_subset_eq_iff simp flip: count_le_replicate_mset_subset_eq) lemma msubseteq_replicate_msetE: assumes "A \# replicate_mset n a" obtains m where "m \ n" and "A = replicate_mset m a" proof (cases "n = 0") case True with assms that show thesis by simp next case False from assms have "set_mset A \ set_mset (replicate_mset n a)" by (rule set_mset_mono) with False have "set_mset A \ {a}" by simp then have "\m. A = replicate_mset m a" proof (induction A) case empty then show ?case by simp next case (add b A) then obtain m where "A = replicate_mset m a" by auto with add.prems show ?case by (auto intro: exI [of _ "Suc m"]) qed then obtain m where A: "A = replicate_mset m a" .. with assms have "m \ n" by (auto simp add: replicate_mset_msubseteq_iff) then show thesis using A .. qed lemma count_image_mset_lt_imp_lt_raw: assumes "finite A" and "A = set_mset M \ set_mset N" and "count (image_mset f M) b < count (image_mset f N) b" shows "\x. f x = b \ count M x < count N x" using assms proof (induct A arbitrary: M N b rule: finite_induct) case (insert x F) note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and cnt_b = this(5) let ?Ma = "{#y \# M. y \ x#}" let ?Mb = "{#y \# M. y = x#}" let ?Na = "{#y \# N. y \ x#}" let ?Nb = "{#y \# N. y = x#}" have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" using multiset_partition by blast+ have f_eq_ma_na: "F = set_mset ?Ma \ set_mset ?Na" using x_f_eq_m_n x_ni_f by auto show ?case proof (cases "count (image_mset f ?Ma) b < count (image_mset f ?Na) b") case cnt_ba: True obtain xa where "f xa = b" and "count ?Ma xa < count ?Na xa" using ih[OF f_eq_ma_na cnt_ba] by blast thus ?thesis by (metis count_filter_mset not_less0) next case cnt_ba: False have fx_eq_b: "f x = b" using cnt_b cnt_ba by (subst (asm) m_part, subst (asm) n_part, auto simp: filter_eq_replicate_mset split: if_splits) moreover have "count M x < count N x" using cnt_b cnt_ba by (subst (asm) m_part, subst (asm) n_part, auto simp: filter_eq_replicate_mset split: if_splits) ultimately show ?thesis by blast qed qed auto lemma count_image_mset_lt_imp_lt: assumes cnt_b: "count (image_mset f M) b < count (image_mset f N) b" shows "\x. f x = b \ count M x < count N x" by (rule count_image_mset_lt_imp_lt_raw[of "set_mset M \ set_mset N", OF _ refl cnt_b]) auto lemma count_image_mset_le_imp_lt_raw: assumes "finite A" and "A = set_mset M \ set_mset N" and "count (image_mset f M) (f a) + count N a < count (image_mset f N) (f a) + count M a" shows "\b. f b = f a \ count M b < count N b" using assms proof (induct A arbitrary: M N rule: finite_induct) case (insert x F) note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and cnt_lt = this(5) let ?Ma = "{#y \# M. y \ x#}" let ?Mb = "{#y \# M. y = x#}" let ?Na = "{#y \# N. y \ x#}" let ?Nb = "{#y \# N. y = x#}" have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" using multiset_partition by blast+ have f_eq_ma_na: "F = set_mset ?Ma \ set_mset ?Na" using x_f_eq_m_n x_ni_f by auto show ?case proof (cases "f x = f a") case fx_ne_fa: False have cnt_fma_fa: "count (image_mset f ?Ma) (f a) = count (image_mset f M) (f a)" using fx_ne_fa by (subst (2) m_part) (auto simp: filter_eq_replicate_mset) have cnt_fna_fa: "count (image_mset f ?Na) (f a) = count (image_mset f N) (f a)" using fx_ne_fa by (subst (2) n_part) (auto simp: filter_eq_replicate_mset) have cnt_ma_a: "count ?Ma a = count M a" using fx_ne_fa by (subst (2) m_part) (auto simp: filter_eq_replicate_mset) have cnt_na_a: "count ?Na a = count N a" using fx_ne_fa by (subst (2) n_part) (auto simp: filter_eq_replicate_mset) obtain b where fb_eq_fa: "f b = f a" and cnt_b: "count ?Ma b < count ?Na b" using ih[OF f_eq_ma_na] cnt_lt unfolding cnt_fma_fa cnt_fna_fa cnt_ma_a cnt_na_a by blast have fx_ne_fb: "f x \ f b" using fb_eq_fa fx_ne_fa by simp have cnt_ma_b: "count ?Ma b = count M b" using fx_ne_fb by (subst (2) m_part) auto have cnt_na_b: "count ?Na b = count N b" using fx_ne_fb by (subst (2) n_part) auto show ?thesis using fb_eq_fa cnt_b unfolding cnt_ma_b cnt_na_b by blast next case fx_eq_fa: True show ?thesis proof (cases "x = a") case x_eq_a: True have "count (image_mset f ?Ma) (f a) + count ?Na a < count (image_mset f ?Na) (f a) + count ?Ma a" using cnt_lt x_eq_a by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, auto simp: filter_eq_replicate_mset) thus ?thesis using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) next case x_ne_a: False show ?thesis proof (cases "count M x < count N x") case True thus ?thesis using fx_eq_fa by blast next case False hence cnt_x: "count M x \ count N x" by fastforce have "count M x + count (image_mset f ?Ma) (f a) + count ?Na a < count N x + count (image_mset f ?Na) (f a) + count ?Ma a" using cnt_lt x_ne_a fx_eq_fa by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, auto simp: filter_eq_replicate_mset) hence "count (image_mset f ?Ma) (f a) + count ?Na a < count (image_mset f ?Na) (f a) + count ?Ma a" using cnt_x by linarith thus ?thesis using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) qed qed qed qed auto lemma count_image_mset_le_imp_lt: assumes "count (image_mset f M) (f a) \ count (image_mset f N) (f a)" and "count M a > count N a" shows "\b. f b = f a \ count M b < count N b" using assms by (auto intro: count_image_mset_le_imp_lt_raw[of "set_mset M \ set_mset N"]) lemma size_filter_unsat_elem: assumes "x \# M" and "\ P x" shows "size {#x \# M. P x#} < size M" proof - have "size (filter_mset P M) \ size M" - using assms by (metis add.right_neutral add_diff_cancel_left' count_filter_mset mem_Collect_eq - multiset_partition nonempty_has_size set_mset_def size_union) + using assms + by (metis dual_order.strict_iff_order filter_mset_eq_conv mset_subset_size subset_mset.nless_le) then show ?thesis by (meson leD nat_neq_iff size_filter_mset_lesseq) qed lemma size_filter_ne_elem: "x \# M \ size {#y \# M. y \ x#} < size M" by (simp add: size_filter_unsat_elem[of x M "\y. y \ x"]) lemma size_eq_ex_count_lt: assumes sz_m_eq_n: "size M = size N" and m_eq_n: "M \ N" shows "\x. count M x < count N x" proof - obtain x where "count M x \ count N x" using m_eq_n by (meson multiset_eqI) moreover { assume "count M x < count N x" hence ?thesis by blast } moreover { assume cnt_x: "count M x > count N x" have "size {#y \# M. y = x#} + size {#y \# M. y \ x#} = size {#y \# N. y = x#} + size {#y \# N. y \ x#}" using sz_m_eq_n multiset_partition by (metis size_union) hence sz_m_minus_x: "size {#y \# M. y \ x#} < size {#y \# N. y \ x#}" using cnt_x by (simp add: filter_eq_replicate_mset) then obtain y where "count {#y \# M. y \ x#} y < count {#y \# N. y \ x#} y" using size_lt_imp_ex_count_lt[OF sz_m_minus_x] by blast hence "count M y < count N y" by (metis count_filter_mset less_asym) hence ?thesis by blast } ultimately show ?thesis by fastforce qed subsection \Big operators\ locale comm_monoid_mset = comm_monoid begin interpretation comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) interpretation comp?: comp_fun_commute "f \ g" by (fact comp_comp_fun_commute) context begin definition F :: "'a multiset \ 'a" where eq_fold: "F M = fold_mset f \<^bold>1 M" lemma empty [simp]: "F {#} = \<^bold>1" by (simp add: eq_fold) lemma singleton [simp]: "F {#x#} = x" proof - interpret comp_fun_commute by standard (simp add: fun_eq_iff left_commute) show ?thesis by (simp add: eq_fold) qed lemma union [simp]: "F (M + N) = F M \<^bold>* F N" proof - interpret comp_fun_commute f by standard (simp add: fun_eq_iff left_commute) show ?thesis by (induct N) (simp_all add: left_commute eq_fold) qed lemma add_mset [simp]: "F (add_mset x N) = x \<^bold>* F N" unfolding add_mset_add_single[of x N] union by (simp add: ac_simps) lemma insert [simp]: shows "F (image_mset g (add_mset x A)) = g x \<^bold>* F (image_mset g A)" by (simp add: eq_fold) lemma remove: assumes "x \# A" shows "F A = x \<^bold>* F (A - {#x#})" using multi_member_split[OF assms] by auto lemma neutral: "\x\#A. x = \<^bold>1 \ F A = \<^bold>1" by (induct A) simp_all lemma neutral_const [simp]: "F (image_mset (\_. \<^bold>1) A) = \<^bold>1" by (simp add: neutral) private lemma F_image_mset_product: "F {#g x j \<^bold>* F {#g i j. i \# A#}. j \# B#} = F (image_mset (g x) B) \<^bold>* F {#F {#g i j. i \# A#}. j \# B#}" by (induction B) (simp_all add: left_commute semigroup.assoc semigroup_axioms) lemma swap: "F (image_mset (\i. F (image_mset (g i) B)) A) = F (image_mset (\j. F (image_mset (\i. g i j) A)) B)" apply (induction A, simp) apply (induction B, auto simp add: F_image_mset_product ac_simps) done lemma distrib: "F (image_mset (\x. g x \<^bold>* h x) A) = F (image_mset g A) \<^bold>* F (image_mset h A)" by (induction A) (auto simp: ac_simps) lemma union_disjoint: "A \# B = {#} \ F (image_mset g (A \# B)) = F (image_mset g A) \<^bold>* F (image_mset g B)" by (induction A) (auto simp: ac_simps) end end lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute ((+) :: 'a multiset \ _ \ _)" by standard (simp add: add_ac comp_def) declare comp_fun_commute.fold_mset_add_mset[OF comp_fun_commute_plus_mset, simp] lemma in_mset_fold_plus_iff[iff]: "x \# fold_mset (+) M NN \ x \# M \ (\N. N \# NN \ x \# N)" by (induct NN) auto context comm_monoid_add begin sublocale sum_mset: comm_monoid_mset plus 0 defines sum_mset = sum_mset.F .. lemma sum_unfold_sum_mset: "sum f A = sum_mset (image_mset f (mset_set A))" by (cases "finite A") (induct A rule: finite_induct, simp_all) end notation sum_mset ("\\<^sub>#") syntax (ASCII) "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10) syntax "_sum_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" ("(3\_\#_. _)" [0, 51, 10] 10) translations "\i \# A. b" \ "CONST sum_mset (CONST image_mset (\i. b) A)" context comm_monoid_add begin lemma sum_mset_sum_list: "sum_mset (mset xs) = sum_list xs" by (induction xs) auto end context canonically_ordered_monoid_add begin lemma sum_mset_0_iff [simp]: "sum_mset M = 0 \ (\x \ set_mset M. x = 0)" by (induction M) auto end context ordered_comm_monoid_add begin lemma sum_mset_mono: "sum_mset (image_mset f K) \ sum_mset (image_mset g K)" if "\i. i \# K \ f i \ g i" using that by (induction K) (simp_all add: add_mono) end context cancel_comm_monoid_add begin lemma sum_mset_diff: "sum_mset (M - N) = sum_mset M - sum_mset N" if "N \# M" for M N :: "'a multiset" using that by (auto simp add: subset_mset.le_iff_add) end context semiring_0 begin lemma sum_mset_distrib_left: "c * (\x \# M. f x) = (\x \# M. c * f(x))" by (induction M) (simp_all add: algebra_simps) lemma sum_mset_distrib_right: "(\x \# M. f x) * c = (\x \# M. f x * c)" by (induction M) (simp_all add: algebra_simps) end lemma sum_mset_product: fixes f :: "'a::{comm_monoid_add,times} \ 'b::semiring_0" shows "(\i \# A. f i) * (\i \# B. g i) = (\i\#A. \j\#B. f i * g j)" by (subst sum_mset.swap) (simp add: sum_mset_distrib_left sum_mset_distrib_right) context semiring_1 begin lemma sum_mset_replicate_mset [simp]: "sum_mset (replicate_mset n a) = of_nat n * a" by (induction n) (simp_all add: algebra_simps) lemma sum_mset_delta: "sum_mset (image_mset (\x. if x = y then c else 0) A) = c * of_nat (count A y)" by (induction A) (simp_all add: algebra_simps) lemma sum_mset_delta': "sum_mset (image_mset (\x. if y = x then c else 0) A) = c * of_nat (count A y)" by (induction A) (simp_all add: algebra_simps) end lemma of_nat_sum_mset [simp]: "of_nat (sum_mset A) = sum_mset (image_mset of_nat A)" by (induction A) auto lemma size_eq_sum_mset: "size M = (\a\#M. 1)" using image_mset_const_eq [of "1::nat" M] by simp lemma size_mset_set [simp]: "size (mset_set A) = card A" by (simp only: size_eq_sum_mset card_eq_sum sum_unfold_sum_mset) lemma sum_mset_constant [simp]: fixes y :: "'b::semiring_1" shows \(\x\#A. y) = of_nat (size A) * y\ by (induction A) (auto simp: algebra_simps) lemma set_mset_Union_mset[simp]: "set_mset (\\<^sub># MM) = (\M \ set_mset MM. set_mset M)" by (induct MM) auto lemma in_Union_mset_iff[iff]: "x \# \\<^sub># MM \ (\M. M \# MM \ x \# M)" by (induct MM) auto lemma count_sum: "count (sum f A) x = sum (\a. count (f a) x) A" by (induct A rule: infinite_finite_induct) simp_all lemma sum_eq_empty_iff: assumes "finite A" shows "sum f A = {#} \ (\a\A. f a = {#})" using assms by induct simp_all lemma Union_mset_empty_conv[simp]: "\\<^sub># M = {#} \ (\i\#M. i = {#})" by (induction M) auto lemma Union_image_single_mset[simp]: "\\<^sub># (image_mset (\x. {#x#}) m) = m" by(induction m) auto lemma size_multiset_sum_mset [simp]: "size (\X\#A. X :: 'a multiset) = (\X\#A. size X)" by (induction A) auto context comm_monoid_mult begin sublocale prod_mset: comm_monoid_mset times 1 defines prod_mset = prod_mset.F .. lemma prod_mset_empty: "prod_mset {#} = 1" by (fact prod_mset.empty) lemma prod_mset_singleton: "prod_mset {#x#} = x" by (fact prod_mset.singleton) lemma prod_mset_Un: "prod_mset (A + B) = prod_mset A * prod_mset B" by (fact prod_mset.union) lemma prod_mset_prod_list: "prod_mset (mset xs) = prod_list xs" by (induct xs) auto lemma prod_mset_replicate_mset [simp]: "prod_mset (replicate_mset n a) = a ^ n" by (induct n) simp_all lemma prod_unfold_prod_mset: "prod f A = prod_mset (image_mset f (mset_set A))" by (cases "finite A") (induct A rule: finite_induct, simp_all) lemma prod_mset_multiplicity: "prod_mset M = prod (\x. x ^ count M x) (set_mset M)" by (simp add: fold_mset_def prod.eq_fold prod_mset.eq_fold funpow_times_power comp_def) lemma prod_mset_delta: "prod_mset (image_mset (\x. if x = y then c else 1) A) = c ^ count A y" by (induction A) simp_all lemma prod_mset_delta': "prod_mset (image_mset (\x. if y = x then c else 1) A) = c ^ count A y" by (induction A) simp_all lemma prod_mset_subset_imp_dvd: assumes "A \# B" shows "prod_mset A dvd prod_mset B" proof - from assms have "B = (B - A) + A" by (simp add: subset_mset.diff_add) also have "prod_mset \ = prod_mset (B - A) * prod_mset A" by simp also have "prod_mset A dvd \" by simp finally show ?thesis . qed lemma dvd_prod_mset: assumes "x \# A" shows "x dvd prod_mset A" using assms prod_mset_subset_imp_dvd [of "{#x#}" A] by simp end notation prod_mset ("\\<^sub>#") syntax (ASCII) "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3PROD _:#_. _)" [0, 51, 10] 10) syntax "_prod_mset_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" ("(3\_\#_. _)" [0, 51, 10] 10) translations "\i \# A. b" \ "CONST prod_mset (CONST image_mset (\i. b) A)" lemma prod_mset_constant [simp]: "(\_\#A. c) = c ^ size A" by (simp add: image_mset_const_eq) lemma (in semidom) prod_mset_zero_iff [iff]: "prod_mset A = 0 \ 0 \# A" by (induct A) auto lemma (in semidom_divide) prod_mset_diff: assumes "B \# A" and "0 \# B" shows "prod_mset (A - B) = prod_mset A div prod_mset B" proof - from assms obtain C where "A = B + C" by (metis subset_mset.add_diff_inverse) with assms show ?thesis by simp qed lemma (in semidom_divide) prod_mset_minus: assumes "a \# A" and "a \ 0" shows "prod_mset (A - {#a#}) = prod_mset A div a" using assms prod_mset_diff [of "{#a#}" A] by auto lemma (in normalization_semidom) normalize_prod_mset_normalize: "normalize (prod_mset (image_mset normalize A)) = normalize (prod_mset A)" proof (induction A) case (add x A) have "normalize (prod_mset (image_mset normalize (add_mset x A))) = normalize (x * normalize (prod_mset (image_mset normalize A)))" by simp also note add.IH finally show ?case by simp qed auto lemma (in algebraic_semidom) is_unit_prod_mset_iff: "is_unit (prod_mset A) \ (\x \# A. is_unit x)" by (induct A) (auto simp: is_unit_mult_iff) lemma (in normalization_semidom_multiplicative) normalize_prod_mset: "normalize (prod_mset A) = prod_mset (image_mset normalize A)" by (induct A) (simp_all add: normalize_mult) lemma (in normalization_semidom_multiplicative) normalized_prod_msetI: assumes "\a. a \# A \ normalize a = a" shows "normalize (prod_mset A) = prod_mset A" proof - from assms have "image_mset normalize A = A" by (induct A) simp_all then show ?thesis by (simp add: normalize_prod_mset) qed lemma image_prod_mset_multiplicity: "prod_mset (image_mset f M) = prod (\x. f x ^ count M x) (set_mset M)" proof (induction M) case (add x M) show ?case proof (cases "x \ set_mset M") case True have "(\y\set_mset (add_mset x M). f y ^ count (add_mset x M) y) = (\y\set_mset M. (if y = x then f x else 1) * f y ^ count M y)" using True add by (intro prod.cong) auto also have "\ = f x * (\y\set_mset M. f y ^ count M y)" using True by (subst prod.distrib) auto also note add.IH [symmetric] finally show ?thesis using True by simp next case False hence "(\y\set_mset (add_mset x M). f y ^ count (add_mset x M) y) = f x * (\y\set_mset M. f y ^ count (add_mset x M) y)" by (auto simp: not_in_iff) also have "(\y\set_mset M. f y ^ count (add_mset x M) y) = (\y\set_mset M. f y ^ count M y)" using False by (intro prod.cong) auto also note add.IH [symmetric] finally show ?thesis by simp qed qed auto subsection \Multiset as order-ignorant lists\ context linorder begin lemma mset_insort [simp]: "mset (insort_key k x xs) = add_mset x (mset xs)" by (induct xs) simp_all lemma mset_sort [simp]: "mset (sort_key k xs) = mset xs" by (induct xs) simp_all text \ This lemma shows which properties suffice to show that a function \f\ with \f xs = ys\ behaves like sort. \ lemma properties_for_sort_key: assumes "mset ys = mset xs" and "\k. k \ set ys \ filter (\x. f k = f x) ys = filter (\x. f k = f x) xs" and "sorted (map f ys)" shows "sort_key f xs = ys" using assms proof (induct xs arbitrary: ys) case Nil then show ?case by simp next case (Cons x xs) from Cons.prems(2) have "\k \ set ys. filter (\x. f k = f x) (remove1 x ys) = filter (\x. f k = f x) xs" by (simp add: filter_remove1) with Cons.prems have "sort_key f xs = remove1 x ys" by (auto intro!: Cons.hyps simp add: sorted_map_remove1) moreover from Cons.prems have "x \# mset ys" by auto then have "x \ set ys" by simp ultimately show ?case using Cons.prems by (simp add: insort_key_remove1) qed lemma properties_for_sort: assumes multiset: "mset ys = mset xs" and "sorted ys" shows "sort xs = ys" proof (rule properties_for_sort_key) from multiset show "mset ys = mset xs" . from \sorted ys\ show "sorted (map (\x. x) ys)" by simp from multiset have "length (filter (\y. k = y) ys) = length (filter (\x. k = x) xs)" for k by (rule mset_eq_length_filter) then have "replicate (length (filter (\y. k = y) ys)) k = replicate (length (filter (\x. k = x) xs)) k" for k by simp then show "k \ set ys \ filter (\y. k = y) ys = filter (\x. k = x) xs" for k by (simp add: replicate_length_filter) qed lemma sort_key_inj_key_eq: assumes mset_equal: "mset xs = mset ys" and "inj_on f (set xs)" and "sorted (map f ys)" shows "sort_key f xs = ys" proof (rule properties_for_sort_key) from mset_equal show "mset ys = mset xs" by simp from \sorted (map f ys)\ show "sorted (map f ys)" . show "[x\ys . f k = f x] = [x\xs . f k = f x]" if "k \ set ys" for k proof - from mset_equal have set_equal: "set xs = set ys" by (rule mset_eq_setD) with that have "insert k (set ys) = set ys" by auto with \inj_on f (set xs)\ have inj: "inj_on f (insert k (set ys))" by (simp add: set_equal) from inj have "[x\ys . f k = f x] = filter (HOL.eq k) ys" by (auto intro!: inj_on_filter_key_eq) also have "\ = replicate (count (mset ys) k) k" by (simp add: replicate_count_mset_eq_filter_eq) also have "\ = replicate (count (mset xs) k) k" using mset_equal by simp also have "\ = filter (HOL.eq k) xs" by (simp add: replicate_count_mset_eq_filter_eq) also have "\ = [x\xs . f k = f x]" using inj by (auto intro!: inj_on_filter_key_eq [symmetric] simp add: set_equal) finally show ?thesis . qed qed lemma sort_key_eq_sort_key: assumes "mset xs = mset ys" and "inj_on f (set xs)" shows "sort_key f xs = sort_key f ys" by (rule sort_key_inj_key_eq) (simp_all add: assms) lemma sort_key_by_quicksort: "sort_key f xs = sort_key f [x\xs. f x < f (xs ! (length xs div 2))] @ [x\xs. f x = f (xs ! (length xs div 2))] @ sort_key f [x\xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs") proof (rule properties_for_sort_key) show "mset ?rhs = mset ?lhs" by (rule multiset_eqI) auto show "sorted (map f ?rhs)" by (auto simp add: sorted_append intro: sorted_map_same) next fix l assume "l \ set ?rhs" let ?pivot = "f (xs ! (length xs div 2))" have *: "\x. f l = f x \ f x = f l" by auto have "[x \ sort_key f xs . f x = f l] = [x \ xs. f x = f l]" unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) with * have **: "[x \ sort_key f xs . f l = f x] = [x \ xs. f l = f x]" by simp have "\x P. P (f x) ?pivot \ f l = f x \ P (f l) ?pivot \ f l = f x" by auto then have "\P. [x \ sort_key f xs . P (f x) ?pivot \ f l = f x] = [x \ sort_key f xs. P (f l) ?pivot \ f l = f x]" by simp note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"] show "[x \ ?rhs. f l = f x] = [x \ ?lhs. f l = f x]" proof (cases "f l" ?pivot rule: linorder_cases) case less then have "f l \ ?pivot" and "\ f l > ?pivot" by auto with less show ?thesis by (simp add: filter_sort [symmetric] ** ***) next case equal then show ?thesis by (simp add: * less_le) next case greater then have "f l \ ?pivot" and "\ f l < ?pivot" by auto with greater show ?thesis by (simp add: filter_sort [symmetric] ** ***) qed qed lemma sort_by_quicksort: "sort xs = sort [x\xs. x < xs ! (length xs div 2)] @ [x\xs. x = xs ! (length xs div 2)] @ sort [x\xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") using sort_key_by_quicksort [of "\x. x", symmetric] by simp text \A stable parameterized quicksort\ definition part :: "('b \ 'a) \ 'a \ 'b list \ 'b list \ 'b list \ 'b list" where "part f pivot xs = ([x \ xs. f x < pivot], [x \ xs. f x = pivot], [x \ xs. pivot < f x])" lemma part_code [code]: "part f pivot [] = ([], [], [])" "part f pivot (x # xs) = (let (lts, eqs, gts) = part f pivot xs; x' = f x in if x' < pivot then (x # lts, eqs, gts) else if x' > pivot then (lts, eqs, x # gts) else (lts, x # eqs, gts))" by (auto simp add: part_def Let_def split_def) lemma sort_key_by_quicksort_code [code]: "sort_key f xs = (case xs of [] \ [] | [x] \ xs | [x, y] \ (if f x \ f y then xs else [y, x]) | _ \ let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs in sort_key f lts @ eqs @ sort_key f gts)" proof (cases xs) case Nil then show ?thesis by simp next case (Cons _ ys) note hyps = Cons show ?thesis proof (cases ys) case Nil with hyps show ?thesis by simp next case (Cons _ zs) note hyps = hyps Cons show ?thesis proof (cases zs) case Nil with hyps show ?thesis by auto next case Cons from sort_key_by_quicksort [of f xs] have "sort_key f xs = (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs in sort_key f lts @ eqs @ sort_key f gts)" by (simp only: split_def Let_def part_def fst_conv snd_conv) with hyps Cons show ?thesis by (simp only: list.cases) qed qed qed end hide_const (open) part lemma mset_remdups_subset_eq: "mset (remdups xs) \# mset xs" by (induct xs) (auto intro: subset_mset.order_trans) lemma mset_update: "i < length ls \ mset (ls[i := v]) = add_mset v (mset ls - {#ls ! i#})" proof (induct ls arbitrary: i) case Nil then show ?case by simp next case (Cons x xs) show ?case proof (cases i) case 0 then show ?thesis by simp next case (Suc i') with Cons show ?thesis by (cases \x = xs ! i'\) auto qed qed lemma mset_swap: "i < length ls \ j < length ls \ mset (ls[j := ls ! i, i := ls ! j]) = mset ls" by (cases "i = j") (simp_all add: mset_update nth_mem_mset) lemma mset_eq_finite: \finite {ys. mset ys = mset xs}\ proof - have \{ys. mset ys = mset xs} \ {ys. set ys \ set xs \ length ys \ length xs}\ by (auto simp add: dest: mset_eq_setD mset_eq_length) moreover have \finite {ys. set ys \ set xs \ length ys \ length xs}\ using finite_lists_length_le by blast ultimately show ?thesis by (rule finite_subset) qed subsection \The multiset order\ definition mult1 :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult1 r = {(N, M). \a M0 K. M = add_mset a M0 \ N = M0 + K \ (\b. b \# K \ (b, a) \ r)}" definition mult :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where "mult r = (mult1 r)\<^sup>+" definition multp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp r M N \ (M, N) \ mult {(x, y). r x y}" declare multp_def[pred_set_conv] lemma mult1I: assumes "M = add_mset a M0" and "N = M0 + K" and "\b. b \# K \ (b, a) \ r" shows "(N, M) \ mult1 r" using assms unfolding mult1_def by blast lemma mult1E: assumes "(N, M) \ mult1 r" obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "\b. b \# K \ (b, a) \ r" using assms unfolding mult1_def by blast lemma mono_mult1: assumes "r \ r'" shows "mult1 r \ mult1 r'" unfolding mult1_def using assms by blast lemma mono_mult: assumes "r \ r'" shows "mult r \ mult r'" unfolding mult_def using mono_mult1[OF assms] trancl_mono by blast lemma mono_multp[mono]: "r \ r' \ multp r \ multp r'" unfolding le_fun_def le_bool_def proof (intro allI impI) fix M N :: "'a multiset" assume "\x xa. r x xa \ r' x xa" hence "{(x, y). r x y} \ {(x, y). r' x y}" by blast thus "multp r M N \ multp r' M N" unfolding multp_def by (fact mono_mult[THEN subsetD, rotated]) qed lemma not_less_empty [iff]: "(M, {#}) \ mult1 r" by (simp add: mult1_def) subsubsection \Well-foundedness\ lemma less_add: assumes mult1: "(N, add_mset a M0) \ mult1 r" shows "(\M. (M, M0) \ mult1 r \ N = add_mset a M) \ (\K. (\b. b \# K \ (b, a) \ r) \ N = M0 + K)" proof - let ?r = "\K a. \b. b \# K \ (b, a) \ r" let ?R = "\N M. \a M0 K. M = add_mset a M0 \ N = M0 + K \ ?r K a" obtain a' M0' K where M0: "add_mset a M0 = add_mset a' M0'" and N: "N = M0' + K" and r: "?r K a'" using mult1 unfolding mult1_def by auto show ?thesis (is "?case1 \ ?case2") proof - from M0 consider "M0 = M0'" "a = a'" | K' where "M0 = add_mset a' K'" "M0' = add_mset a K'" by atomize_elim (simp only: add_eq_conv_ex) then show ?thesis proof cases case 1 with N r have "?r K a \ N = M0 + K" by simp then have ?case2 .. then show ?thesis .. next case 2 from N 2(2) have n: "N = add_mset a (K' + K)" by simp with r 2(1) have "?R (K' + K) M0" by blast with n have ?case1 by (simp add: mult1_def) then show ?thesis .. qed qed qed lemma all_accessible: assumes "wf r" shows "\M. M \ Wellfounded.acc (mult1 r)" proof let ?R = "mult1 r" let ?W = "Wellfounded.acc ?R" { fix M M0 a assume M0: "M0 \ ?W" and wf_hyp: "\b. (b, a) \ r \ (\M \ ?W. add_mset b M \ ?W)" and acc_hyp: "\M. (M, M0) \ ?R \ add_mset a M \ ?W" have "add_mset a M0 \ ?W" proof (rule accI [of "add_mset a M0"]) fix N assume "(N, add_mset a M0) \ ?R" then consider M where "(M, M0) \ ?R" "N = add_mset a M" | K where "\b. b \# K \ (b, a) \ r" "N = M0 + K" by atomize_elim (rule less_add) then show "N \ ?W" proof cases case 1 from acc_hyp have "(M, M0) \ ?R \ add_mset a M \ ?W" .. from this and \(M, M0) \ ?R\ have "add_mset a M \ ?W" .. then show "N \ ?W" by (simp only: \N = add_mset a M\) next case 2 from this(1) have "M0 + K \ ?W" proof (induct K) case empty from M0 show "M0 + {#} \ ?W" by simp next case (add x K) from add.prems have "(x, a) \ r" by simp with wf_hyp have "\M \ ?W. add_mset x M \ ?W" by blast moreover from add have "M0 + K \ ?W" by simp ultimately have "add_mset x (M0 + K) \ ?W" .. then show "M0 + (add_mset x K) \ ?W" by simp qed then show "N \ ?W" by (simp only: 2(2)) qed qed } note tedious_reasoning = this show "M \ ?W" for M proof (induct M) show "{#} \ ?W" proof (rule accI) fix b assume "(b, {#}) \ ?R" with not_less_empty show "b \ ?W" by contradiction qed fix M a assume "M \ ?W" from \wf r\ have "\M \ ?W. add_mset a M \ ?W" proof induct fix a assume r: "\b. (b, a) \ r \ (\M \ ?W. add_mset b M \ ?W)" show "\M \ ?W. add_mset a M \ ?W" proof fix M assume "M \ ?W" then show "add_mset a M \ ?W" by (rule acc_induct) (rule tedious_reasoning [OF _ r]) qed qed from this and \M \ ?W\ show "add_mset a M \ ?W" .. qed qed lemma wf_mult1: "wf r \ wf (mult1 r)" by (rule acc_wfI) (rule all_accessible) lemma wf_mult: "wf r \ wf (mult r)" unfolding mult_def by (rule wf_trancl) (rule wf_mult1) lemma wfP_multp: "wfP r \ wfP (multp r)" unfolding multp_def wfP_def by (simp add: wf_mult) subsubsection \Closure-free presentation\ text \One direction.\ lemma mult_implies_one_step: assumes trans: "trans r" and MN: "(M, N) \ mult r" shows "\I J K. N = I + J \ M = I + K \ J \ {#} \ (\k \ set_mset K. \j \ set_mset J. (k, j) \ r)" using MN unfolding mult_def mult1_def proof (induction rule: converse_trancl_induct) case (base y) then show ?case by force next case (step y z) note yz = this(1) and zN = this(2) and N_decomp = this(3) obtain I J K where N: "N = I + J" "z = I + K" "J \ {#}" "\k\#K. \j\#J. (k, j) \ r" using N_decomp by blast obtain a M0 K' where z: "z = add_mset a M0" and y: "y = M0 + K'" and K: "\b. b \# K' \ (b, a) \ r" using yz by blast show ?case proof (cases "a \# K") case True moreover have "\j\#J. (k, j) \ r" if "k \# K'" for k using K N trans True by (meson that transE) ultimately show ?thesis by (rule_tac x = I in exI, rule_tac x = J in exI, rule_tac x = "(K - {#a#}) + K'" in exI) (use z y N in \auto simp del: subset_mset.add_diff_assoc2 dest: in_diffD\) next case False then have "a \# I" by (metis N(2) union_iff union_single_eq_member z) moreover have "M0 = I + K - {#a#}" using N(2) z by force ultimately show ?thesis by (rule_tac x = "I - {#a#}" in exI, rule_tac x = "add_mset a J" in exI, rule_tac x = "K + K'" in exI) (use z y N False K in \auto simp: add.assoc\) qed qed lemma multp_implies_one_step: "transp R \ multp R M N \ \I J K. N = I + J \ M = I + K \ J \ {#} \ (\k\#K. \x\#J. R k x)" by (rule mult_implies_one_step[to_pred]) lemma one_step_implies_mult: assumes "J \ {#}" and "\k \ set_mset K. \j \ set_mset J. (k, j) \ r" shows "(I + K, I + J) \ mult r" using assms proof (induction "size J" arbitrary: I J K) case 0 then show ?case by auto next case (Suc n) note IH = this(1) and size_J = this(2)[THEN sym] obtain J' a where J: "J = add_mset a J'" using size_J by (blast dest: size_eq_Suc_imp_eq_union) show ?case proof (cases "J' = {#}") case True then show ?thesis using J Suc by (fastforce simp add: mult_def mult1_def) next case [simp]: False have K: "K = {#x \# K. (x, a) \ r#} + {#x \# K. (x, a) \ r#}" by simp have "(I + K, (I + {# x \# K. (x, a) \ r #}) + J') \ mult r" using IH[of J' "{# x \# K. (x, a) \ r#}" "I + {# x \# K. (x, a) \ r#}"] J Suc.prems K size_J by (auto simp: ac_simps) moreover have "(I + {#x \# K. (x, a) \ r#} + J', I + J) \ mult r" by (fastforce simp: J mult1_def mult_def) ultimately show ?thesis unfolding mult_def by simp qed qed lemma one_step_implies_multp: "J \ {#} \ \k\#K. \j\#J. R k j \ multp R (I + K) (I + J)" by (rule one_step_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def, simplified]) lemma subset_implies_mult: assumes sub: "A \# B" shows "(A, B) \ mult r" proof - have ApBmA: "A + (B - A) = B" using sub by simp have BmA: "B - A \ {#}" using sub by (simp add: Diff_eq_empty_iff_mset subset_mset.less_le_not_le) thus ?thesis by (rule one_step_implies_mult[of "B - A" "{#}" _ A, unfolded ApBmA, simplified]) qed lemma subset_implies_multp: "A \# B \ multp r A B" by (rule subset_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def]) lemma multp_repeat_mset_repeat_msetI: assumes "transp R" and "multp R A B" and "n \ 0" shows "multp R (repeat_mset n A) (repeat_mset n B)" proof - from \transp R\ \multp R A B\ obtain I J K where "B = I + J" and "A = I + K" and "J \ {#}" and "\k \# K. \x \# J. R k x" by (auto dest: multp_implies_one_step) have repeat_n_A_eq: "repeat_mset n A = repeat_mset n I + repeat_mset n K" using \A = I + K\ by simp have repeat_n_B_eq: "repeat_mset n B = repeat_mset n I + repeat_mset n J" using \B = I + J\ by simp show ?thesis unfolding repeat_n_A_eq repeat_n_B_eq proof (rule one_step_implies_multp) from \n \ 0\ show "repeat_mset n J \ {#}" using \J \ {#}\ by (simp add: repeat_mset_eq_empty_iff) next show "\k \# repeat_mset n K. \j \# repeat_mset n J. R k j" using \\k \# K. \x \# J. R k x\ by (metis count_greater_zero_iff nat_0_less_mult_iff repeat_mset.rep_eq) qed qed subsubsection \Monotonicity\ lemma multp_mono_strong: assumes "multp R M1 M2" and "transp R" and S_if_R: "\x y. x \ set_mset M1 \ y \ set_mset M2 \ R x y \ S x y" shows "multp S M1 M2" proof - obtain I J K where "M2 = I + J" and "M1 = I + K" and "J \ {#}" and "\k\#K. \x\#J. R k x" using multp_implies_one_step[OF \transp R\ \multp R M1 M2\] by auto show ?thesis unfolding \M2 = I + J\ \M1 = I + K\ proof (rule one_step_implies_multp[OF \J \ {#}\]) show "\k\#K. \j\#J. S k j" using S_if_R by (metis \M1 = I + K\ \M2 = I + J\ \\k\#K. \x\#J. R k x\ union_iff) qed qed lemma mult_mono_strong: assumes "(M1, M2) \ mult r" and "trans r" and S_if_R: "\x y. x \ set_mset M1 \ y \ set_mset M2 \ (x, y) \ r \ (x, y) \ s" shows "(M1, M2) \ mult s" using assms multp_mono_strong[of "\x y. (x, y) \ r" M1 M2 "\x y. (x, y) \ s", unfolded multp_def transp_trans_eq, simplified] by blast lemma monotone_on_multp_multp_image_mset: assumes "monotone_on A orda ordb f" and "transp orda" shows "monotone_on {M. set_mset M \ A} (multp orda) (multp ordb) (image_mset f)" proof (rule monotone_onI) fix M1 M2 assume M1_in: "M1 \ {M. set_mset M \ A}" and M2_in: "M2 \ {M. set_mset M \ A}" and M1_lt_M2: "multp orda M1 M2" from multp_implies_one_step[OF \transp orda\ M1_lt_M2] obtain I J K where M2_eq: "M2 = I + J" and M1_eq: "M1 = I + K" and J_neq_mempty: "J \ {#}" and ball_K_less: "\k\#K. \x\#J. orda k x" by metis have "multp ordb (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)" proof (intro one_step_implies_multp ballI) show "image_mset f J \ {#}" using J_neq_mempty by simp next fix k' assume "k'\#image_mset f K" then obtain k where "k' = f k" and k_in: "k \# K" by auto then obtain j where j_in: "j\#J" and "orda k j" using ball_K_less by auto have "ordb (f k) (f j)" proof (rule \monotone_on A orda ordb f\[THEN monotone_onD, OF _ _ \orda k j\]) show "k \ A" using M1_eq M1_in k_in by auto next show "j \ A" using M2_eq M2_in j_in by auto qed thus "\j\#image_mset f J. ordb k' j" using \j \# J\ \k' = f k\ by auto qed thus "multp ordb (image_mset f M1) (image_mset f M2)" by (simp add: M1_eq M2_eq) qed lemma monotone_multp_multp_image_mset: assumes "monotone orda ordb f" and "transp orda" shows "monotone (multp orda) (multp ordb) (image_mset f)" by (rule monotone_on_multp_multp_image_mset[OF assms, simplified]) lemma multp_image_mset_image_msetI: assumes "multp (\x y. R (f x) (f y)) M1 M2" and "transp R" shows "multp R (image_mset f M1) (image_mset f M2)" proof - from \transp R\ have "transp (\x y. R (f x) (f y))" by (auto intro: transpI dest: transpD) with \multp (\x y. R (f x) (f y)) M1 M2\ obtain I J K where "M2 = I + J" and "M1 = I + K" and "J \ {#}" and "\k\#K. \x\#J. R (f k) (f x)" using multp_implies_one_step by blast have "multp R (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)" proof (rule one_step_implies_multp) show "image_mset f J \ {#}" by (simp add: \J \ {#}\) next show "\k\#image_mset f K. \j\#image_mset f J. R k j" by (simp add: \\k\#K. \x\#J. R (f k) (f x)\) qed thus ?thesis by (simp add: \M1 = I + K\ \M2 = I + J\) qed lemma multp_image_mset_image_msetD: assumes "multp R (image_mset f A) (image_mset f B)" and "transp R" and inj_on_f: "inj_on f (set_mset A \ set_mset B)" shows "multp (\x y. R (f x) (f y)) A B" proof - from assms(1,2) obtain I J K where f_B_eq: "image_mset f B = I + J" and f_A_eq: "image_mset f A = I + K" and J_neq_mempty: "J \ {#}" and ball_K_less: "\k\#K. \x\#J. R k x" by (auto dest: multp_implies_one_step) from f_B_eq obtain I' J' where B_def: "B = I' + J'" and I_def: "I = image_mset f I'" and J_def: "J = image_mset f J'" using image_mset_eq_plusD by blast from inj_on_f have inj_on_f': "inj_on f (set_mset A \ set_mset I')" by (rule inj_on_subset) (auto simp add: B_def) from f_A_eq obtain K' where A_def: "A = I' + K'" and K_def: "K = image_mset f K'" by (auto simp: I_def dest: image_mset_eq_image_mset_plusD[OF _ inj_on_f']) show ?thesis unfolding A_def B_def proof (intro one_step_implies_multp ballI) from J_neq_mempty show "J' \ {#}" by (simp add: J_def) next fix k assume "k \# K'" with ball_K_less obtain j' where "j' \# J" and "R (f k) j'" using K_def by auto moreover then obtain j where "j \# J'" and "f j = j'" using J_def by auto ultimately show "\j\#J'. R (f k) (f j)" by blast qed qed subsubsection \The multiset extension is cancellative for multiset union\ lemma mult_cancel: assumes "trans s" and "irrefl_on (set_mset Z) s" shows "(X + Z, Y + Z) \ mult s \ (X, Y) \ mult s" (is "?L \ ?R") proof assume ?L thus ?R using \irrefl_on (set_mset Z) s\ proof (induct Z) case (add z Z) obtain X' Y' Z' where *: "add_mset z X + Z = Z' + X'" "add_mset z Y + Z = Z' + Y'" "Y' \ {#}" "\x \ set_mset X'. \y \ set_mset Y'. (x, y) \ s" using mult_implies_one_step[OF \trans s\ add(2)] by auto consider Z2 where "Z' = add_mset z Z2" | X2 Y2 where "X' = add_mset z X2" "Y' = add_mset z Y2" using *(1,2) by (metis add_mset_remove_trivial_If insert_iff set_mset_add_mset_insert union_iff) thus ?case proof (cases) case 1 thus ?thesis using * one_step_implies_mult[of Y' X' s Z2] add(3) by (auto simp: add.commute[of _ "{#_#}"] add.assoc intro: add(1) elim: irrefl_on_subset) next case 2 then obtain y where "y \ set_mset Y2" "(z, y) \ s" using *(4) \irrefl_on (set_mset (add_mset z Z)) s\ by (auto simp: irrefl_on_def) moreover from this transD[OF \trans s\ _ this(2)] have "x' \ set_mset X2 \ \y \ set_mset Y2. (x', y) \ s" for x' using 2 *(4)[rule_format, of x'] by auto ultimately show ?thesis using * one_step_implies_mult[of Y2 X2 s Z'] 2 add(3) by (force simp: add.commute[of "{#_#}"] add.assoc[symmetric] intro: add(1) elim: irrefl_on_subset) qed qed auto next assume ?R then obtain I J K where "Y = I + J" "X = I + K" "J \ {#}" "\k \ set_mset K. \j \ set_mset J. (k, j) \ s" using mult_implies_one_step[OF \trans s\] by blast thus ?L using one_step_implies_mult[of J K s "I + Z"] by (auto simp: ac_simps) qed lemma multp_cancel: "transp R \ irreflp_on (set_mset Z) R \ multp R (X + Z) (Y + Z) \ multp R X Y" by (rule mult_cancel[to_pred]) lemma mult_cancel_add_mset: "trans r \ irrefl_on {z} r \ ((add_mset z X, add_mset z Y) \ mult r) = ((X, Y) \ mult r)" by (rule mult_cancel[of _ "{#_#}", simplified]) lemma multp_cancel_add_mset: "transp R \ irreflp_on {z} R \ multp R (add_mset z X) (add_mset z Y) = multp R X Y" by (rule mult_cancel_add_mset[to_pred, folded bot_set_def]) lemma mult_cancel_max0: assumes "trans s" and "irrefl_on (set_mset X \ set_mset Y) s" shows "(X, Y) \ mult s \ (X - X \# Y, Y - X \# Y) \ mult s" (is "?L \ ?R") proof - have "(X - X \# Y + X \# Y, Y - X \# Y + X \# Y) \ mult s \ (X - X \# Y, Y - X \# Y) \ mult s" proof (rule mult_cancel) from assms show "trans s" by simp next from assms show "irrefl_on (set_mset (X \# Y)) s" by simp qed moreover have "X - X \# Y + X \# Y = X" "Y - X \# Y + X \# Y = Y" by (auto simp flip: count_inject) ultimately show ?thesis by simp qed lemma mult_cancel_max: "trans r \ irrefl_on (set_mset X \ set_mset Y) r \ (X, Y) \ mult r \ (X - Y, Y - X) \ mult r" by (rule mult_cancel_max0[simplified]) lemma multp_cancel_max: "transp R \ irreflp_on (set_mset X \ set_mset Y) R \ multp R X Y \ multp R (X - Y) (Y - X)" by (rule mult_cancel_max[to_pred]) subsubsection \Strict partial-order properties\ lemma mult1_lessE: assumes "(N, M) \ mult1 {(a, b). r a b}" and "asymp r" obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "a \# K" "\b. b \# K \ r b a" proof - from assms obtain a M0 K where "M = add_mset a M0" "N = M0 + K" and *: "b \# K \ r b a" for b by (blast elim: mult1E) moreover from * [of a] have "a \# K" using \asymp r\ by (meson asympD) ultimately show thesis by (auto intro: that) qed lemma trans_on_mult: assumes "trans_on A r" and "\M. M \ B \ set_mset M \ A" shows "trans_on B (mult r)" using assms by (metis mult_def subset_UNIV trans_on_subset trans_trancl) lemma trans_mult: "trans r \ trans (mult r)" using trans_on_mult[of UNIV r UNIV, simplified] . lemma transp_on_multp: assumes "transp_on A r" and "\M. M \ B \ set_mset M \ A" shows "transp_on B (multp r)" by (metis mult_def multp_def transD trans_trancl transp_onI) lemma transp_multp: "transp r \ transp (multp r)" using transp_on_multp[of UNIV r UNIV, simplified] . lemma irrefl_mult: assumes "trans r" "irrefl r" shows "irrefl (mult r)" proof (intro irreflI notI) fix M assume "(M, M) \ mult r" then obtain I J K where "M = I + J" and "M = I + K" and "J \ {#}" and "(\k\set_mset K. \j\set_mset J. (k, j) \ r)" using mult_implies_one_step[OF \trans r\] by blast then have *: "K \ {#}" and **: "\k\set_mset K. \j\set_mset K. (k, j) \ r" by auto have "finite (set_mset K)" by simp hence "set_mset K = {}" using ** proof (induction rule: finite_induct) case empty thus ?case by simp next case (insert x F) have False using \irrefl r\[unfolded irrefl_def, rule_format] using \trans r\[THEN transD] by (metis equals0D insert.IH insert.prems insertE insertI1 insertI2) thus ?case .. qed with * show False by simp qed lemma irreflp_multp: "transp R \ irreflp R \ irreflp (multp R)" by (rule irrefl_mult[of "{(x, y). r x y}" for r, folded transp_trans_eq irreflp_irrefl_eq, simplified, folded multp_def]) instantiation multiset :: (preorder) order begin definition less_multiset :: "'a multiset \ 'a multiset \ bool" where "M < N \ multp (<) M N" definition less_eq_multiset :: "'a multiset \ 'a multiset \ bool" where "less_eq_multiset M N \ M < N \ M = N" instance proof intro_classes fix M N :: "'a multiset" show "(M < N) = (M \ N \ \ N \ M)" unfolding less_eq_multiset_def less_multiset_def by (metis irreflp_def irreflp_on_less irreflp_multp transpE transp_on_less transp_multp) next fix M :: "'a multiset" show "M \ M" unfolding less_eq_multiset_def by simp next fix M1 M2 M3 :: "'a multiset" show "M1 \ M2 \ M2 \ M3 \ M1 \ M3" unfolding less_eq_multiset_def less_multiset_def using transp_multp[OF transp_on_less, THEN transpD] by blast next fix M N :: "'a multiset" show "M \ N \ N \ M \ M = N" unfolding less_eq_multiset_def less_multiset_def using transp_multp[OF transp_on_less, THEN transpD] using irreflp_multp[OF transp_on_less irreflp_on_less, unfolded irreflp_def, rule_format] by blast qed end lemma mset_le_irrefl [elim!]: fixes M :: "'a::preorder multiset" shows "M < M \ R" by simp lemma wfP_less_multiset[simp]: assumes wfP_less: "wfP ((<) :: ('a :: preorder) \ 'a \ bool)" shows "wfP ((<) :: 'a multiset \ 'a multiset \ bool)" unfolding less_multiset_def using wfP_multp[OF wfP_less] . subsubsection \Strict total-order properties\ lemma total_on_mult: assumes "total_on A r" and "trans r" and "\M. M \ B \ set_mset M \ A" shows "total_on B (mult r)" proof (rule total_onI) fix M1 M2 assume "M1 \ B" and "M2 \ B" and "M1 \ M2" let ?I = "M1 \# M2" show "(M1, M2) \ mult r \ (M2, M1) \ mult r" proof (cases "M1 - ?I = {#} \ M2 - ?I = {#}") case True with \M1 \ M2\ show ?thesis by (metis Diff_eq_empty_iff_mset diff_intersect_left_idem diff_intersect_right_idem subset_implies_mult subset_mset.less_le) next case False from assms(1) have "total_on (set_mset (M1 - ?I)) r" by (meson \M1 \ B\ assms(3) diff_subset_eq_self set_mset_mono total_on_subset) with False obtain greatest1 where greatest1_in: "greatest1 \# M1 - ?I" and greatest1_greatest: "\x \# M1 - ?I. greatest1 \ x \ (x, greatest1) \ r" using Multiset.bex_greatest_element[to_set, of "M1 - ?I" r] by (metis assms(2) subset_UNIV trans_on_subset) from assms(1) have "total_on (set_mset (M2 - ?I)) r" by (meson \M2 \ B\ assms(3) diff_subset_eq_self set_mset_mono total_on_subset) with False obtain greatest2 where greatest2_in: "greatest2 \# M2 - ?I" and greatest2_greatest: "\x \# M2 - ?I. greatest2 \ x \ (x, greatest2) \ r" using Multiset.bex_greatest_element[to_set, of "M2 - ?I" r] by (metis assms(2) subset_UNIV trans_on_subset) have "greatest1 \ greatest2" using greatest1_in \greatest2 \# M2 - ?I\ by (metis diff_intersect_left_idem diff_intersect_right_idem dual_order.eq_iff in_diff_count in_diff_countE le_add_same_cancel2 less_irrefl zero_le) hence "(greatest1, greatest2) \ r \ (greatest2, greatest1) \ r" using \total_on A r\[unfolded total_on_def, rule_format, of greatest1 greatest2] \M1 \ B\ \M2 \ B\ greatest1_in greatest2_in assms(3) by (meson in_diffD in_mono) thus ?thesis proof (elim disjE) assume "(greatest1, greatest2) \ r" have "(?I + (M1 - ?I), ?I + (M2 - ?I)) \ mult r" proof (rule one_step_implies_mult[of "M2 - ?I" "M1 - ?I" r ?I]) show "M2 - ?I \ {#}" using False by force next show "\k\#M1 - ?I. \j\#M2 - ?I. (k, j) \ r" using \(greatest1, greatest2) \ r\ greatest2_in greatest1_greatest by (metis assms(2) transD) qed hence "(M1, M2) \ mult r" by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1 subset_mset.inf.cobounded2) thus "(M1, M2) \ mult r \ (M2, M1) \ mult r" .. next assume "(greatest2, greatest1) \ r" have "(?I + (M2 - ?I), ?I + (M1 - ?I)) \ mult r" proof (rule one_step_implies_mult[of "M1 - ?I" "M2 - ?I" r ?I]) show "M1 - M1 \# M2 \ {#}" using False by force next show "\k\#M2 - ?I. \j\#M1 - ?I. (k, j) \ r" using \(greatest2, greatest1) \ r\ greatest1_in greatest2_greatest by (metis assms(2) transD) qed hence "(M2, M1) \ mult r" by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1 subset_mset.inf.cobounded2) thus "(M1, M2) \ mult r \ (M2, M1) \ mult r" .. qed qed qed lemma total_mult: "total r \ trans r \ total (mult r)" by (rule total_on_mult[of UNIV r UNIV, simplified]) lemma totalp_on_multp: "totalp_on A R \ transp R \ (\M. M \ B \ set_mset M \ A) \ totalp_on B (multp R)" using total_on_mult[of A "{(x,y). R x y}" B, to_pred] by (simp add: multp_def total_on_def totalp_on_def) lemma totalp_multp: "totalp R \ transp R \ totalp (multp R)" by (rule totalp_on_multp[of UNIV R UNIV, simplified]) subsection \Quasi-executable version of the multiset extension\ text \ Predicate variants of \mult\ and the reflexive closure of \mult\, which are executable whenever the given predicate \P\ is. Together with the standard code equations for \(\#\) and \(-\) this should yield quadratic (with respect to calls to \P\) implementations of \multp_code\ and \multeqp_code\. \ definition multp_code :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp_code P N M = (let Z = M \# N; X = M - Z in X \ {#} \ (let Y = N - Z in (\y \ set_mset Y. \x \ set_mset X. P y x)))" definition multeqp_code :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multeqp_code P N M = (let Z = M \# N; X = M - Z; Y = N - Z in (\y \ set_mset Y. \x \ set_mset X. P y x))" lemma multp_code_iff_mult: assumes "irrefl_on (set_mset N \ set_mset M) R" and "trans R" and [simp]: "\x y. P x y \ (x, y) \ R" shows "multp_code P N M \ (N, M) \ mult R" (is "?L \ ?R") proof - have *: "M \# N + (N - M \# N) = N" "M \# N + (M - M \# N) = M" "(M - M \# N) \# (N - M \# N) = {#}" by (auto simp flip: count_inject) show ?thesis proof assume ?L thus ?R using one_step_implies_mult[of "M - M \# N" "N - M \# N" R "M \# N"] * by (auto simp: multp_code_def Let_def) next { fix I J K :: "'a multiset" assume "(I + J) \# (I + K) = {#}" then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty) } note [dest!] = this assume ?R thus ?L using mult_cancel_max using mult_implies_one_step[OF assms(2), of "N - M \# N" "M - M \# N"] mult_cancel_max[OF assms(2,1)] * by (auto simp: multp_code_def) qed qed lemma multp_code_iff_multp: "irreflp_on (set_mset M \ set_mset N) R \ transp R \ multp_code R M N \ multp R M N" using multp_code_iff_mult[simplified, to_pred, of M N R R] by simp lemma multp_code_eq_multp: assumes "irreflp R" and "transp R" shows "multp_code R = multp R" proof (intro ext) fix M N show "multp_code R M N = multp R M N" proof (rule multp_code_iff_multp) from assms show "irreflp_on (set_mset M \ set_mset N) R" by (auto intro: irreflp_on_subset) next from assms show "transp R" by simp qed qed lemma multeqp_code_iff_reflcl_mult: assumes "irrefl_on (set_mset N \ set_mset M) R" and "trans R" and "\x y. P x y \ (x, y) \ R" shows "multeqp_code P N M \ (N, M) \ (mult R)\<^sup>=" proof - { assume "N \ M" "M - M \# N = {#}" then obtain y where "count N y \ count M y" by (auto simp flip: count_inject) then have "\y. count M y < count N y" using \M - M \# N = {#}\ by (auto simp flip: count_inject dest!: le_neq_implies_less fun_cong[of _ _ y]) } then have "multeqp_code P N M \ multp_code P N M \ N = M" by (auto simp: multeqp_code_def multp_code_def Let_def in_diff_count) thus ?thesis using multp_code_iff_mult[OF assms] by simp qed lemma multeqp_code_iff_reflclp_multp: "irreflp_on (set_mset M \ set_mset N) R \ transp R \ multeqp_code R M N \ (multp R)\<^sup>=\<^sup>= M N" using multeqp_code_iff_reflcl_mult[simplified, to_pred, of M N R R] by simp lemma multeqp_code_eq_reflclp_multp: assumes "irreflp R" and "transp R" shows "multeqp_code R = (multp R)\<^sup>=\<^sup>=" proof (intro ext) fix M N show "multeqp_code R M N \ (multp R)\<^sup>=\<^sup>= M N" proof (rule multeqp_code_iff_reflclp_multp) from assms show "irreflp_on (set_mset M \ set_mset N) R" by (auto intro: irreflp_on_subset) next from assms show "transp R" by simp qed qed subsubsection \Monotonicity of multiset union\ lemma mult1_union: "(B, D) \ mult1 r \ (C + B, C + D) \ mult1 r" by (force simp: mult1_def) lemma union_le_mono2: "B < D \ C + B < C + (D::'a::preorder multiset)" -apply (unfold less_multiset_def multp_def mult_def) -apply (erule trancl_induct) - apply (blast intro: mult1_union) -apply (blast intro: mult1_union trancl_trans) -done + unfolding less_multiset_def multp_def mult_def + by (induction rule: trancl_induct; blast intro: mult1_union trancl_trans) lemma union_le_mono1: "B < D \ B + C < D + (C::'a::preorder multiset)" -apply (subst add.commute [of B C]) -apply (subst add.commute [of D C]) -apply (erule union_le_mono2) -done + by (metis add.commute union_le_mono2) lemma union_less_mono: fixes A B C D :: "'a::preorder multiset" shows "A < C \ B < D \ A + B < C + D" by (blast intro!: union_le_mono1 union_le_mono2 less_trans) instantiation multiset :: (preorder) ordered_ab_semigroup_add begin instance by standard (auto simp add: less_eq_multiset_def intro: union_le_mono2) end subsubsection \Termination proofs with multiset orders\ lemma multi_member_skip: "x \# XS \ x \# {# y #} + XS" and multi_member_this: "x \# {# x #} + XS" and multi_member_last: "x \# {# x #}" by auto definition "ms_strict = mult pair_less" definition "ms_weak = ms_strict \ Id" lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)" -unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def -by (auto intro: wf_mult1 wf_trancl simp: mult_def) + unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def + by (auto intro: wf_mult1 wf_trancl simp: mult_def) lemma smsI: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z + B) \ ms_strict" unfolding ms_strict_def by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases) lemma wmsI: "(set_mset A, set_mset B) \ max_strict \ A = {#} \ B = {#} \ (Z + A, Z + B) \ ms_weak" unfolding ms_weak_def ms_strict_def by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult) inductive pw_leq where pw_leq_empty: "pw_leq {#} {#}" | pw_leq_step: "\(x,y) \ pair_leq; pw_leq X Y \ \ pw_leq ({#x#} + X) ({#y#} + Y)" lemma pw_leq_lstep: "(x, y) \ pair_leq \ pw_leq {#x#} {#y#}" by (drule pw_leq_step) (rule pw_leq_empty, simp) lemma pw_leq_split: assumes "pw_leq X Y" shows "\A B Z. X = A + Z \ Y = B + Z \ ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" using assms proof induct case pw_leq_empty thus ?case by auto next case (pw_leq_step x y X Y) then obtain A B Z where [simp]: "X = A + Z" "Y = B + Z" and 1[simp]: "(set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#})" by auto from pw_leq_step consider "x = y" | "(x, y) \ pair_less" unfolding pair_leq_def by auto thus ?case proof cases case [simp]: 1 have "{#x#} + X = A + ({#y#}+Z) \ {#y#} + Y = B + ({#y#}+Z) \ ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" by auto thus ?thesis by blast next case 2 let ?A' = "{#x#} + A" and ?B' = "{#y#} + B" have "{#x#} + X = ?A' + Z" "{#y#} + Y = ?B' + Z" by auto moreover have "(set_mset ?A', set_mset ?B') \ max_strict" using 1 2 unfolding max_strict_def by (auto elim!: max_ext.cases) ultimately show ?thesis by blast qed qed lemma assumes pwleq: "pw_leq Z Z'" shows ms_strictI: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_strict" and ms_weakI1: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_weak" and ms_weakI2: "(Z + {#}, Z' + {#}) \ ms_weak" proof - from pw_leq_split[OF pwleq] obtain A' B' Z'' where [simp]: "Z = A' + Z''" "Z' = B' + Z''" and mx_or_empty: "(set_mset A', set_mset B') \ max_strict \ (A' = {#} \ B' = {#})" by blast { assume max: "(set_mset A, set_mset B) \ max_strict" from mx_or_empty have "(Z'' + (A + A'), Z'' + (B + B')) \ ms_strict" proof assume max': "(set_mset A', set_mset B') \ max_strict" with max have "(set_mset (A + A'), set_mset (B + B')) \ max_strict" by (auto simp: max_strict_def intro: max_ext_additive) thus ?thesis by (rule smsI) next assume [simp]: "A' = {#} \ B' = {#}" show ?thesis by (rule smsI) (auto intro: max) qed thus "(Z + A, Z' + B) \ ms_strict" by (simp add: ac_simps) thus "(Z + A, Z' + B) \ ms_weak" by (simp add: ms_weak_def) } from mx_or_empty have "(Z'' + A', Z'' + B') \ ms_weak" by (rule wmsI) thus "(Z + {#}, Z' + {#}) \ ms_weak" by (simp add: ac_simps) qed lemma empty_neutral: "{#} + x = x" "x + {#} = x" and nonempty_plus: "{# x #} + rs \ {#}" and nonempty_single: "{# x #} \ {#}" by auto setup \ let fun msetT T = \<^Type>\multiset T\; fun mk_mset T [] = \<^instantiate>\'a = T in term \{#}\\ | mk_mset T [x] = \<^instantiate>\'a = T and x in term \{#x#}\\ | mk_mset T (x :: xs) = \<^Const>\plus \msetT T\ for \mk_mset T [x]\ \mk_mset T xs\\ fun mset_member_tac ctxt m i = if m <= 0 then resolve_tac ctxt @{thms multi_member_this} i ORELSE resolve_tac ctxt @{thms multi_member_last} i else resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i fun mset_nonempty_tac ctxt = resolve_tac ctxt @{thms nonempty_plus} ORELSE' resolve_tac ctxt @{thms nonempty_single} fun regroup_munion_conv ctxt = Function_Lib.regroup_conv ctxt \<^const_abbrev>\empty_mset\ \<^const_name>\plus\ (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral})) fun unfold_pwleq_tac ctxt i = (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st)) ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i) ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i) val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union}, @{thm Un_insert_left}, @{thm Un_empty_left}] in ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset { msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps, smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1}, reduction_pair = @{thm ms_reduction_pair} }) end \ subsection \Legacy theorem bindings\ lemmas multi_count_eq = multiset_eq_iff [symmetric] lemma union_commute: "M + N = N + (M::'a multiset)" by (fact add.commute) lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" by (fact add.assoc) lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" by (fact add.left_commute) lemmas union_ac = union_assoc union_commute union_lcomm add_mset_commute lemma union_right_cancel: "M + K = N + K \ M = (N::'a multiset)" by (fact add_right_cancel) lemma union_left_cancel: "K + M = K + N \ M = (N::'a multiset)" by (fact add_left_cancel) lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \ X = Y" by (fact add_left_imp_eq) lemma mset_subset_trans: "(M::'a multiset) \# K \ K \# N \ M \# N" by (fact subset_mset.less_trans) lemma multiset_inter_commute: "A \# B = B \# A" by (fact subset_mset.inf.commute) lemma multiset_inter_assoc: "A \# (B \# C) = A \# B \# C" by (fact subset_mset.inf.assoc [symmetric]) lemma multiset_inter_left_commute: "A \# (B \# C) = B \# (A \# C)" by (fact subset_mset.inf.left_commute) lemmas multiset_inter_ac = multiset_inter_commute multiset_inter_assoc multiset_inter_left_commute lemma mset_le_not_refl: "\ M < (M::'a::preorder multiset)" by (fact less_irrefl) lemma mset_le_trans: "K < M \ M < N \ K < (N::'a::preorder multiset)" by (fact less_trans) lemma mset_le_not_sym: "M < N \ \ N < (M::'a::preorder multiset)" by (fact less_not_sym) lemma mset_le_asym: "M < N \ (\ P \ N < (M::'a::preorder multiset)) \ P" by (fact less_asym) declaration \ let fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') = let val (maybe_opt, ps) = Nitpick_Model.dest_plain_fun t' ||> (~~) ||> map (apsnd (snd o HOLogic.dest_number)) fun elems_for t = (case AList.lookup (=) ps t of SOME n => replicate n t | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]) in (case maps elems_for (all_values elem_T) @ (if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of [] => \<^Const>\Groups.zero T\ | ts => foldl1 (fn (s, t) => \<^Const>\add_mset elem_T for s t\) ts) end | multiset_postproc _ _ _ _ t = t in Nitpick_Model.register_term_postprocessor \<^typ>\'a multiset\ multiset_postproc end \ subsection \Naive implementation using lists\ code_datatype mset lemma [code]: "{#} = mset []" by simp lemma [code]: "add_mset x (mset xs) = mset (x # xs)" by simp lemma [code]: "Multiset.is_empty (mset xs) \ List.null xs" by (simp add: Multiset.is_empty_def List.null_def) lemma union_code [code]: "mset xs + mset ys = mset (xs @ ys)" by simp lemma [code]: "image_mset f (mset xs) = mset (map f xs)" by simp lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)" by simp lemma [code]: "mset xs - mset ys = mset (fold remove1 ys xs)" by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute diff_diff_add) lemma [code]: "mset xs \# mset ys = mset (snd (fold (\x (ys, zs). if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, [])))" proof - have "\zs. mset (snd (fold (\x (ys, zs). if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) = (mset xs \# mset ys) + mset zs" by (induct xs arbitrary: ys) (auto simp add: inter_add_right1 inter_add_right2 ac_simps) then show ?thesis by simp qed lemma [code]: "mset xs \# mset ys = mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" proof - have "\zs. mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) = (mset xs \# mset ys) + mset zs" by (induct xs arbitrary: ys) (simp_all add: multiset_eq_iff) then show ?thesis by simp qed declare in_multiset_in_set [code_unfold] lemma [code]: "count (mset xs) x = fold (\y. if x = y then Suc else id) xs 0" proof - have "\n. fold (\y. if x = y then Suc else id) xs n = count (mset xs) x + n" by (induct xs) simp_all then show ?thesis by simp qed declare set_mset_mset [code] declare sorted_list_of_multiset_mset [code] lemma [code]: \ \not very efficient, but representation-ignorant!\ "mset_set A = mset (sorted_list_of_set A)" - apply (cases "finite A") - apply simp_all - apply (induct A rule: finite_induct) - apply simp_all - done + by (metis mset_sorted_list_of_multiset sorted_list_of_mset_set) declare size_mset [code] fun subset_eq_mset_impl :: "'a list \ 'a list \ bool option" where "subset_eq_mset_impl [] ys = Some (ys \ [])" | "subset_eq_mset_impl (Cons x xs) ys = (case List.extract ((=) x) ys of None \ None | Some (ys1,_,ys2) \ subset_eq_mset_impl xs (ys1 @ ys2))" lemma subset_eq_mset_impl: "(subset_eq_mset_impl xs ys = None \ \ mset xs \# mset ys) \ (subset_eq_mset_impl xs ys = Some True \ mset xs \# mset ys) \ (subset_eq_mset_impl xs ys = Some False \ mset xs = mset ys)" proof (induct xs arbitrary: ys) case (Nil ys) show ?case by (auto simp: subset_mset.zero_less_iff_neq_zero) next case (Cons x xs ys) show ?case proof (cases "List.extract ((=) x) ys") case None hence x: "x \ set ys" by (simp add: extract_None_iff) { assume "mset (x # xs) \# mset ys" from set_mset_mono[OF this] x have False by simp } note nle = this moreover { assume "mset (x # xs) \# mset ys" hence "mset (x # xs) \# mset ys" by auto from nle[OF this] have False . } ultimately show ?thesis using None by auto next case (Some res) obtain ys1 y ys2 where res: "res = (ys1,y,ys2)" by (cases res, auto) note Some = Some[unfolded res] from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp hence id: "mset ys = add_mset x (mset (ys1 @ ys2))" by auto show ?thesis unfolding subset_eq_mset_impl.simps - unfolding Some option.simps split - unfolding id - using Cons[of "ys1 @ ys2"] - unfolding subset_mset_def subseteq_mset_def by auto + by (simp add: Some id Cons) qed qed lemma [code]: "mset xs \# mset ys \ subset_eq_mset_impl xs ys \ None" - using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) + by (simp add: subset_eq_mset_impl) lemma [code]: "mset xs \# mset ys \ subset_eq_mset_impl xs ys = Some True" - using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) + using subset_eq_mset_impl by blast instantiation multiset :: (equal) equal begin definition [code del]: "HOL.equal A (B :: 'a multiset) \ A = B" lemma [code]: "HOL.equal (mset xs) (mset ys) \ subset_eq_mset_impl xs ys = Some False" unfolding equal_multiset_def using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) instance by standard (simp add: equal_multiset_def) end declare sum_mset_sum_list [code] lemma [code]: "prod_mset (mset xs) = fold times xs 1" proof - have "\x. fold times xs x = prod_mset (mset xs) * x" by (induct xs) (simp_all add: ac_simps) then show ?thesis by simp qed text \ Exercise for the casual reader: add implementations for \<^term>\(\)\ and \<^term>\(<)\ (multiset order). \ text \Quickcheck generators\ context includes term_syntax begin definition msetify :: "'a::typerep list \ (unit \ Code_Evaluation.term) \ 'a multiset \ (unit \ Code_Evaluation.term)" where [code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\} xs" end instantiation multiset :: (random) random begin context includes state_combinator_syntax begin definition "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\xs. Pair (msetify xs))" instance .. end end instantiation multiset :: (full_exhaustive) full_exhaustive begin definition full_exhaustive_multiset :: "('a multiset \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" where "full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (\xs. f (msetify xs)) i" instance .. end hide_const (open) msetify subsection \BNF setup\ definition rel_mset where "rel_mset R X Y \ (\xs ys. mset xs = X \ mset ys = Y \ list_all2 R xs ys)" lemma mset_zip_take_Cons_drop_twice: assumes "length xs = length ys" "j \ length xs" shows "mset (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) = add_mset (x,y) (mset (zip xs ys))" using assms proof (induct xs ys arbitrary: x y j rule: list_induct2) case Nil thus ?case by simp next case (Cons x xs y ys) thus ?case proof (cases "j = 0") case True thus ?thesis by simp next case False then obtain k where k: "j = Suc k" by (cases j) simp hence "k \ length xs" using Cons.prems by auto hence "mset (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) = add_mset (x,y) (mset (zip xs ys))" by (rule Cons.hyps(2)) thus ?thesis unfolding k by auto qed qed lemma ex_mset_zip_left: assumes "length xs = length ys" "mset xs' = mset xs" shows "\ys'. length ys' = length xs' \ mset (zip xs' ys') = mset (zip xs ys)" using assms proof (induct xs ys arbitrary: xs' rule: list_induct2) case Nil thus ?case by auto next case (Cons x xs y ys xs') obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) define xsa where "xsa = take j xs' @ drop (Suc j) xs'" have "mset xs' = {#x#} + mset xsa" unfolding xsa_def using j_len nth_j by (metis Cons_nth_drop_Suc union_mset_add_mset_right add_mset_remove_trivial add_diff_cancel_left' append_take_drop_id mset.simps(2) mset_append) hence ms_x: "mset xsa = mset xs" by (simp add: Cons.prems) then obtain ysa where len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" using Cons.hyps(2) by blast define ys' where "ys' = take j ysa @ y # drop j ysa" have xs': "xs' = take j xsa @ x # drop j xsa" using ms_x j_len nth_j Cons.prems xsa_def by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons length_drop size_mset) have j_len': "j \ length xsa" using j_len xs' xsa_def by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) have "length ys' = length xs'" unfolding ys'_def using Cons.prems len_a ms_x by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" unfolding xs' ys'_def by (rule trans[OF mset_zip_take_Cons_drop_twice]) (auto simp: len_a ms_a j_len') ultimately show ?case by blast qed lemma list_all2_reorder_left_invariance: assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs" shows "\ys'. list_all2 R xs' ys' \ mset ys' = mset ys" proof - have len: "length xs = length ys" using rel list_all2_conv_all_nth by auto obtain ys' where len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" using len ms_x by (metis ex_mset_zip_left) have "list_all2 R xs' ys'" using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) moreover have "mset ys' = mset ys" using len len' ms_xy map_snd_zip mset_map by metis ultimately show ?thesis by blast qed lemma ex_mset: "\xs. mset xs = X" by (induct X) (simp, metis mset.simps(2)) inductive pred_mset :: "('a \ bool) \ 'a multiset \ bool" where "pred_mset P {#}" | "\P a; pred_mset P M\ \ pred_mset P (add_mset a M)" lemma pred_mset_iff: \ \TODO: alias for \<^const>\Multiset.Ball\\ \pred_mset P M \ Multiset.Ball M P\ (is \?P \ ?Q\) proof assume ?P then show ?Q by induction simp_all next assume ?Q then show ?P by (induction M) (auto intro: pred_mset.intros) qed bnf "'a multiset" map: image_mset sets: set_mset bd: natLeq wits: "{#}" rel: rel_mset pred: pred_mset proof - - show "image_mset id = id" - by (rule image_mset.id) show "image_mset (g \ f) = image_mset g \ image_mset f" for f g unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality) show "(\z. z \ set_mset X \ f z = g z) \ image_mset f X = image_mset g X" for f g X by (induct X) simp_all - show "set_mset \ image_mset f = (`) f \ set_mset" for f - by auto show "card_order natLeq" by (rule natLeq_card_order) show "BNF_Cardinal_Arithmetic.cinfinite natLeq" by (rule natLeq_cinfinite) show "regularCard natLeq" by (rule regularCard_natLeq) show "ordLess2 (card_of (set_mset X)) natLeq" for X by transfer (auto simp: finite_iff_ordLess_natLeq[symmetric]) show "rel_mset R OO rel_mset S \ rel_mset (R OO S)" for R S unfolding rel_mset_def[abs_def] OO_def - apply clarify - subgoal for X Z Y xs ys' ys zs - apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys]) - apply (auto intro: list_all2_trans) - done - done + by (smt (verit, ccfv_SIG) list_all2_reorder_left_invariance list_all2_trans predicate2I) show "rel_mset R = (\x y. \z. set_mset z \ {(x, y). R x y} \ image_mset fst z = x \ image_mset snd z = y)" for R unfolding rel_mset_def[abs_def] - apply (rule ext)+ - apply safe - apply (rule_tac x = "mset (zip xs ys)" in exI; - auto simp: in_set_zip list_all2_iff simp flip: mset_map) - apply (rename_tac XY) - apply (cut_tac X = XY in ex_mset) - apply (erule exE) - apply (rename_tac xys) - apply (rule_tac x = "map fst xys" in exI) - apply (auto simp: mset_map) - apply (rule_tac x = "map snd xys" in exI) - apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd) - done - show "z \ set_mset {#} \ False" for z - by auto + by (metis (no_types, lifting) ex_mset list.in_rel mem_Collect_eq mset_map set_mset_mset) show "pred_mset P = (\x. Ball (set_mset x) P)" for P by (simp add: fun_eq_iff pred_mset_iff) -qed +qed auto inductive rel_mset' :: \('a \ 'b \ bool) \ 'a multiset \ 'b multiset \ bool\ where Zero[intro]: "rel_mset' R {#} {#}" | Plus[intro]: "\R a b; rel_mset' R M N\ \ rel_mset' R (add_mset a M) (add_mset b N)" lemma rel_mset_Zero: "rel_mset R {#} {#}" unfolding rel_mset_def Grp_def by auto declare multiset.count[simp] declare count_Abs_multiset[simp] declare multiset.count_inverse[simp] lemma rel_mset_Plus: assumes ab: "R a b" and MN: "rel_mset R M N" shows "rel_mset R (add_mset a M) (add_mset b N)" proof - have "\ya. add_mset a (image_mset fst y) = image_mset fst ya \ add_mset b (image_mset snd y) = image_mset snd ya \ set_mset ya \ {(x, y). R x y}" if "R a b" and "set_mset y \ {(x, y). R x y}" for y using that by (intro exI[of _ "add_mset (a,b) y"]) auto thus ?thesis using assms unfolding multiset.rel_compp_Grp Grp_def by blast qed lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \ rel_mset R M N" by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus) lemma rel_mset_size: "rel_mset R M N \ size M = size N" unfolding multiset.rel_compp_Grp Grp_def by auto lemma rel_mset_Zero_iff [simp]: shows "rel_mset rel {#} Y \ Y = {#}" and "rel_mset rel X {#} \ X = {#}" by (auto simp add: rel_mset_Zero dest: rel_mset_size) lemma multiset_induct2[case_names empty addL addR]: assumes empty: "P {#} {#}" and addL: "\a M N. P M N \ P (add_mset a M) N" and addR: "\a M N. P M N \ P M (add_mset a N)" shows "P M N" -apply(induct N rule: multiset_induct) - apply(induct M rule: multiset_induct, rule empty, erule addL) - apply(induct M rule: multiset_induct, erule addR, erule addR) -done + by (induct N rule: multiset_induct; induct M rule: multiset_induct) (auto simp: assms) lemma multiset_induct2_size[consumes 1, case_names empty add]: assumes c: "size M = size N" and empty: "P {#} {#}" and add: "\a b M N a b. P M N \ P (add_mset a M) (add_mset b N)" shows "P M N" using c proof (induct M arbitrary: N rule: measure_induct_rule[of size]) case (less M) show ?case proof(cases "M = {#}") case True hence "N = {#}" using less.prems by auto thus ?thesis using True empty by auto next case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) have "N \ {#}" using False less.prems by auto then obtain N1 b where N: "N = add_mset b N1" by (metis multi_nonempty_split) have "size M1 = size N1" using less.prems unfolding M N by auto thus ?thesis using M N less.hyps add by auto qed qed lemma msed_map_invL: assumes "image_mset f (add_mset a M) = N" shows "\N1. N = add_mset (f a) N1 \ image_mset f M = N1" proof - have "f a \# N" using assms multiset.set_map[of f "add_mset a M"] by auto then obtain N1 where N: "N = add_mset (f a) N1" using multi_member_split by metis have "image_mset f M = N1" using assms unfolding N by simp thus ?thesis using N by blast qed lemma msed_map_invR: assumes "image_mset f M = add_mset b N" shows "\M1 a. M = add_mset a M1 \ f a = b \ image_mset f M1 = N" proof - obtain a where a: "a \# M" and fa: "f a = b" using multiset.set_map[of f M] unfolding assms by (metis image_iff union_single_eq_member) then obtain M1 where M: "M = add_mset a M1" using multi_member_split by metis have "image_mset f M1 = N" using assms unfolding M fa[symmetric] by simp thus ?thesis using M fa by blast qed lemma msed_rel_invL: assumes "rel_mset R (add_mset a M) N" shows "\N1 b. N = add_mset b N1 \ R a b \ rel_mset R M N1" proof - obtain K where KM: "image_mset fst K = add_mset a M" and KN: "image_mset snd K = N" and sK: "set_mset K \ {(a, b). R a b}" using assms unfolding multiset.rel_compp_Grp Grp_def by auto obtain K1 ab where K: "K = add_mset ab K1" and a: "fst ab = a" and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto obtain N1 where N: "N = add_mset (snd ab) N1" and K1N1: "image_mset snd K1 = N1" using msed_map_invL[OF KN[unfolded K]] by auto have Rab: "R a (snd ab)" using sK a unfolding K by auto have "rel_mset R M N1" using sK K1M K1N1 unfolding K multiset.rel_compp_Grp Grp_def by auto thus ?thesis using N Rab by auto qed lemma msed_rel_invR: assumes "rel_mset R M (add_mset b N)" shows "\M1 a. M = add_mset a M1 \ R a b \ rel_mset R M1 N" proof - obtain K where KN: "image_mset snd K = add_mset b N" and KM: "image_mset fst K = M" and sK: "set_mset K \ {(a, b). R a b}" using assms unfolding multiset.rel_compp_Grp Grp_def by auto obtain K1 ab where K: "K = add_mset ab K1" and b: "snd ab = b" and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto obtain M1 where M: "M = add_mset (fst ab) M1" and K1M1: "image_mset fst K1 = M1" using msed_map_invL[OF KM[unfolded K]] by auto have Rab: "R (fst ab) b" using sK b unfolding K by auto have "rel_mset R M1 N" using sK K1N K1M1 unfolding K multiset.rel_compp_Grp Grp_def by auto thus ?thesis using M Rab by auto qed lemma rel_mset_imp_rel_mset': assumes "rel_mset R M N" shows "rel_mset' R M N" using assms proof(induct M arbitrary: N rule: measure_induct_rule[of size]) case (less M) have c: "size M = size N" using rel_mset_size[OF less.prems] . show ?case proof(cases "M = {#}") case True hence "N = {#}" using c by simp thus ?thesis using True rel_mset'.Zero by auto next case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) obtain N1 b where N: "N = add_mset b N1" and R: "R a b" and ms: "rel_mset R M1 N1" using msed_rel_invL[OF less.prems[unfolded M]] by auto have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp thus ?thesis using rel_mset'.Plus[of R a b, OF R] unfolding M N by simp qed qed lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N" using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto text \The main end product for \<^const>\rel_mset\: inductive characterization:\ lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] = rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]] subsection \Size setup\ lemma size_multiset_o_map: "size_multiset g \ image_mset f = size_multiset (g \ f)" apply (rule ext) subgoal for x by (induct x) auto done setup \ BNF_LFP_Size.register_size_global \<^type_name>\multiset\ \<^const_name>\size_multiset\ @{thm size_multiset_overloaded_def} @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single size_union} @{thms size_multiset_o_map} \ subsection \Lemmas about Size\ lemma size_mset_SucE: "size A = Suc n \ (\a B. A = {#a#} + B \ size B = n \ P) \ P" by (cases A) (auto simp add: ac_simps) lemma size_Suc_Diff1: "x \# M \ Suc (size (M - {#x#})) = size M" using arg_cong[OF insert_DiffM, of _ _ size] by simp lemma size_Diff_singleton: "x \# M \ size (M - {#x#}) = size M - 1" by (simp flip: size_Suc_Diff1) lemma size_Diff_singleton_if: "size (A - {#x#}) = (if x \# A then size A - 1 else size A)" by (simp add: diff_single_trivial size_Diff_singleton) lemma size_Un_Int: "size A + size B = size (A \# B) + size (A \# B)" by (metis inter_subset_eq_union size_union subset_mset.diff_add union_diff_inter_eq_sup) lemma size_Un_disjoint: "A \# B = {#} \ size (A \# B) = size A + size B" using size_Un_Int[of A B] by simp lemma size_Diff_subset_Int: "size (M - M') = size M - size (M \# M')" by (metis diff_intersect_left_idem size_Diff_submset subset_mset.inf_le1) lemma diff_size_le_size_Diff: "size (M :: _ multiset) - size M' \ size (M - M')" by (simp add: diff_le_mono2 size_Diff_subset_Int size_mset_mono) lemma size_Diff1_less: "x\# M \ size (M - {#x#}) < size M" by (rule Suc_less_SucD) (simp add: size_Suc_Diff1) lemma size_Diff2_less: "x\# M \ y\# M \ size (M - {#x#} - {#y#}) < size M" by (metis less_imp_diff_less size_Diff1_less size_Diff_subset_Int) lemma size_Diff1_le: "size (M - {#x#}) \ size M" by (cases "x \# M") (simp_all add: size_Diff1_less less_imp_le diff_single_trivial) lemma size_psubset: "M \# M' \ size M < size M' \ M \# M'" using less_irrefl subset_mset_def by blast lifting_update multiset.lifting lifting_forget multiset.lifting hide_const (open) wcount end diff --git a/src/HOL/Library/Poly_Mapping.thy b/src/HOL/Library/Poly_Mapping.thy --- a/src/HOL/Library/Poly_Mapping.thy +++ b/src/HOL/Library/Poly_Mapping.thy @@ -1,1878 +1,1849 @@ (* Author: Andreas Lochbihler, ETH Zurich Author: Florian Haftmann, TU Muenchen with some material ported from HOL Light by LCP *) section \Polynomial mapping: combination of almost everywhere zero functions with an algebraic view\ theory Poly_Mapping imports Groups_Big_Fun Fun_Lexorder More_List begin subsection \Preliminary: auxiliary operations for \emph{almost everywhere zero}\ text \ A central notion for polynomials are functions being \emph{almost everywhere zero}. For these we provide some auxiliary definitions and lemmas. \ lemma finite_mult_not_eq_zero_leftI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. g a * f a \ 0}" proof - have "{a. g a * f a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_rightI: fixes f :: "'b \ 'a :: mult_zero" assumes "finite {a. f a \ 0}" shows "finite {a. f a * g a \ 0}" proof - have "{a. f a * g a \ 0} \ {a. f a \ 0}" by auto then show ?thesis using assms by (rule finite_subset) qed lemma finite_mult_not_eq_zero_prodI: fixes f g :: "'a \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {(a, b). f a * g b \ 0}" proof - from assms have "finite (?F \ ?G)" by blast then have "finite {(a, b). f a \ 0 \ g b \ 0}" by simp then show ?thesis by (rule rev_finite_subset) auto qed lemma finite_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" (is "finite ?F") assumes "finite {b. g b \ 0}" (is "finite ?G") shows "finite {a + b | a b. f a \ 0 \ g b \ 0}" (is "finite ?FG") proof - from assms have "finite (?F \ ?G)" by (simp add: finite_cartesian_product_iff) then have "finite (case_prod plus ` (?F \ ?G))" by (rule finite_imageI) also have "case_prod plus ` (?F \ ?G) = ?FG" by auto finally show ?thesis by simp qed lemma finite_mult_not_eq_zero_sumI: fixes f g :: "'a::monoid_add \ 'b::semiring_0" assumes "finite {a. f a \ 0}" assumes "finite {b. g b \ 0}" shows "finite {a + b | a b. f a * g b \ 0}" proof - from assms have "finite {a + b | a b. f a \ 0 \ g b \ 0}" by (rule finite_not_eq_zero_sumI) then show ?thesis by (rule rev_finite_subset) (auto dest: mult_not_zero) qed lemma finite_Sum_any_not_eq_zero_weakenI: assumes "finite {a. \b. f a b \ 0}" shows "finite {a. Sum_any (f a) \ 0}" proof - have "{a. Sum_any (f a) \ 0} \ {a. \b. f a b \ 0}" by (auto elim: Sum_any.not_neutral_obtains_not_neutral) then show ?thesis using assms by (rule finite_subset) qed context zero begin definition "when" :: "'a \ bool \ 'a" (infixl "when" 20) where "(a when P) = (if P then a else 0)" text \ Case distinctions always complicate matters, particularly when nested. The @{const "when"} operation allows to minimise these if @{term 0} is the false-case value and makes proof obligations much more readable. \ lemma "when" [simp]: "P \ (a when P) = a" "\ P \ (a when P) = 0" by (simp_all add: when_def) lemma when_simps [simp]: "(a when True) = a" "(a when False) = 0" by simp_all lemma when_cong: assumes "P \ Q" and "Q \ a = b" shows "(a when P) = (b when Q)" using assms by (simp add: when_def) lemma zero_when [simp]: "(0 when P) = 0" by (simp add: when_def) lemma when_when: "(a when P when Q) = (a when P \ Q)" by (cases Q) simp_all lemma when_commute: "(a when Q when P) = (a when P when Q)" by (simp add: when_when conj_commute) lemma when_neq_zero [simp]: "(a when P) \ 0 \ P \ a \ 0" by (cases P) simp_all end context monoid_add begin lemma when_add_distrib: "(a + b when P) = (a when P) + (b when P)" by (simp add: when_def) end context semiring_1 begin lemma zero_power_eq: "0 ^ n = (1 when n = 0)" by (simp add: power_0_left) end context comm_monoid_add begin lemma Sum_any_when_equal [simp]: "(\a. (f a when a = b)) = f b" by (simp add: when_def) lemma Sum_any_when_equal' [simp]: "(\a. (f a when b = a)) = f b" by (simp add: when_def) lemma Sum_any_when_independent: "(\a. g a when P) = ((\a. g a) when P)" by (cases P) simp_all lemma Sum_any_when_dependent_prod_right: "(\(a, b). g a when b = h a) = (\a. g a)" proof - have "inj_on (\a. (a, h a)) {a. g a \ 0}" by (rule inj_onI) auto then show ?thesis unfolding Sum_any.expand_set by (rule sum.reindex_cong) auto qed lemma Sum_any_when_dependent_prod_left: "(\(a, b). g b when a = h b) = (\b. g b)" proof - have "(\(a, b). g b when a = h b) = (\(b, a). g b when a = h b)" by (rule Sum_any.reindex_cong [of prod.swap]) (simp_all add: fun_eq_iff) then show ?thesis by (simp add: Sum_any_when_dependent_prod_right) qed end context cancel_comm_monoid_add begin lemma when_diff_distrib: "(a - b when P) = (a when P) - (b when P)" by (simp add: when_def) end context group_add begin lemma when_uminus_distrib: "(- a when P) = - (a when P)" by (simp add: when_def) end context mult_zero begin lemma mult_when: "a * (b when P) = (a * b when P)" by (cases P) simp_all lemma when_mult: "(a when P) * b = (a * b when P)" by (cases P) simp_all end subsection \Type definition\ text \ The following type is of central importance: \ typedef (overloaded) ('a, 'b) poly_mapping ("(_ \\<^sub>0 /_)" [1, 0] 0) = "{f :: 'a \ 'b::zero. finite {x. f x \ 0}}" morphisms lookup Abs_poly_mapping proof - have "(\_::'a. (0 :: 'b)) \ ?poly_mapping" by simp then show ?thesis by (blast intro!: exI) qed declare lookup_inverse [simp] declare lookup_inject [simp] lemma lookup_Abs_poly_mapping [simp]: "finite {x. f x \ 0} \ lookup (Abs_poly_mapping f) = f" using Abs_poly_mapping_inverse [of f] by simp lemma finite_lookup [simp]: "finite {k. lookup f k \ 0}" using poly_mapping.lookup [of f] by simp lemma finite_lookup_nat [simp]: fixes f :: "'a \\<^sub>0 nat" shows "finite {k. 0 < lookup f k}" using poly_mapping.lookup [of f] by simp lemma poly_mapping_eqI: assumes "\k. lookup f k = lookup g k" shows "f = g" using assms unfolding poly_mapping.lookup_inject [symmetric] by blast lemma poly_mapping_eq_iff: "a = b \ lookup a = lookup b" by auto text \ We model the universe of functions being \emph{almost everywhere zero} by means of a separate type @{typ "('a, 'b) poly_mapping"}. For convenience we provide a suggestive infix syntax which is a variant of the usual function space syntax. Conversion between both types happens through the morphisms \begin{quote} @{term_type lookup} \end{quote} \begin{quote} @{term_type Abs_poly_mapping} \end{quote} satisfying \begin{quote} @{thm lookup_inverse} \end{quote} \begin{quote} @{thm lookup_Abs_poly_mapping} \end{quote} Luckily, we have rarely to deal with those low-level morphisms explicitly but rely on Isabelle's \emph{lifting} package with its method \transfer\ and its specification tool \lift_definition\. \ setup_lifting type_definition_poly_mapping code_datatype Abs_poly_mapping\\FIXME? workaround for preventing \code_abstype\ setup\ text \ @{typ "'a \\<^sub>0 'b"} serves distinctive purposes: \begin{enumerate} \item A clever nesting as @{typ "(nat \\<^sub>0 nat) \\<^sub>0 'a"} later in theory \MPoly\ gives a suitable representation type for polynomials \emph{almost for free}: Interpreting @{typ "nat \\<^sub>0 nat"} as a mapping from variable identifiers to exponents yields monomials, and the whole type maps monomials to coefficients. Lets call this the \emph{ultimate interpretation}. \item A further more specialised type isomorphic to @{typ "nat \\<^sub>0 'a"} is apt to direct implementation using code generation \<^cite>\"Haftmann-Nipkow:2010:code"\. \end{enumerate} Note that despite the names \emph{mapping} and \emph{lookup} suggest something implementation-near, it is best to keep @{typ "'a \\<^sub>0 'b"} as an abstract \emph{algebraic} type providing operations like \emph{addition}, \emph{multiplication} without any notion of key-order, data structures etc. This implementations-specific notions are easily introduced later for particular implementations but do not provide any gain for specifying logical properties of polynomials. \ subsection \Additive structure\ text \ The additive structure covers the usual operations \0\, \+\ and (unary and binary) \-\. Recalling the ultimate interpretation, it is obvious that these have just lift the corresponding operations on values to mappings. Isabelle has a rich hierarchy of algebraic type classes, and in such situations of pointwise lifting a typical pattern is to have instantiations for a considerable number of type classes. The operations themselves are specified using \lift_definition\, where the proofs of the \emph{almost everywhere zero} property can be significantly involved. The @{const lookup} operation is supposed to be usable explicitly (unless in other situations where the morphisms between types are somehow internal to the \emph{lifting} package). Hence it is good style to provide explicit rewrite rules how @{const lookup} acts on operations immediately. \ instantiation poly_mapping :: (type, zero) zero begin lift_definition zero_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 0" by simp instance .. end lemma Abs_poly_mapping [simp]: "Abs_poly_mapping (\k. 0) = 0" by (simp add: zero_poly_mapping.abs_eq) lemma lookup_zero [simp]: "lookup 0 k = 0" by transfer rule instantiation poly_mapping :: (type, monoid_add) monoid_add begin lift_definition plus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k + f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x + f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x + f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_add: "lookup (f + g) k = lookup f k + lookup g k" by transfer rule instance poly_mapping :: (type, comm_monoid_add) comm_monoid_add by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ lemma lookup_sum: "lookup (sum pp X) i = sum (\x. lookup (pp x) i) X" by (induction rule: infinite_finite_induct) (auto simp: lookup_add) (*instance poly_mapping :: (type, "{monoid_add, cancel_semigroup_add}") cancel_semigroup_add by intro_classes (transfer, simp add: fun_eq_iff)+*) instantiation poly_mapping :: (type, cancel_comm_monoid_add) cancel_comm_monoid_add begin lift_definition minus_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is "\f1 f2 k. f1 k - f2 k" proof - fix f1 f2 :: "'a \ 'b" assume "finite {k. f1 k \ 0}" and "finite {k. f2 k \ 0}" then have "finite ({k. f1 k \ 0} \ {k. f2 k \ 0})" by auto moreover have "{x. f1 x - f2 x \ 0} \ {k. f1 k \ 0} \ {k. f2 k \ 0}" by auto ultimately show "finite {x. f1 x - f2 x \ 0}" by (blast intro: finite_subset) qed instance by intro_classes (transfer, simp add: fun_eq_iff diff_diff_add)+ end instantiation poly_mapping :: (type, ab_group_add) ab_group_add begin lift_definition uminus_poly_mapping :: "('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is uminus by simp instance by intro_classes (transfer, simp add: fun_eq_iff ac_simps)+ end lemma lookup_uminus [simp]: "lookup (- f) k = - lookup f k" by transfer simp lemma lookup_minus: "lookup (f - g) k = lookup f k - lookup g k" by transfer rule subsection \Multiplicative structure\ instantiation poly_mapping :: (zero, zero_neq_one) zero_neq_one begin lift_definition one_poly_mapping :: "'a \\<^sub>0 'b" is "\k. 1 when k = 0" by simp instance by intro_classes (transfer, simp add: fun_eq_iff) end lemma lookup_one: "lookup 1 k = (1 when k = 0)" by transfer rule lemma lookup_one_zero [simp]: "lookup 1 0 = 1" by transfer simp definition prod_fun :: "('a \ 'b) \ ('a \ 'b) \ 'a::monoid_add \ 'b::semiring_0" where "prod_fun f1 f2 k = (\l. f1 l * (\q. (f2 q when k = l + q)))" lemma prod_fun_unfold_prod: fixes f g :: "'a :: monoid_add \ 'b::semiring_0" assumes fin_f: "finite {a. f a \ 0}" assumes fin_g: "finite {b. g b \ 0}" shows "prod_fun f g k = (\(a, b). f a * g b when k = a + b)" proof - let ?C = "{a. f a \ 0} \ {b. g b \ 0}" from fin_f fin_g have "finite ?C" by blast moreover have "{a. \b. (f a * g b when k = a + b) \ 0} \ {b. \a. (f a * g b when k = a + b) \ 0} \ {a. f a \ 0} \ {b. g b \ 0}" by auto ultimately show ?thesis using fin_g - by (auto simp add: prod_fun_def + by (auto simp: prod_fun_def Sum_any.cartesian_product [of "{a. f a \ 0} \ {b. g b \ 0}"] Sum_any_right_distrib mult_when) qed lemma finite_prod_fun: fixes f1 f2 :: "'a :: monoid_add \ 'b :: semiring_0" assumes fin1: "finite {l. f1 l \ 0}" and fin2: "finite {q. f2 q \ 0}" shows "finite {k. prod_fun f1 f2 k \ 0}" proof - have *: "finite {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" using assms by simp { fix k l have "{q. (f2 q when k = l + q) \ 0} \ {q. f2 q \ 0 \ k = l + q}" by auto with fin2 have "sum f2 {q. f2 q \ 0 \ k = l + q} = (\q. (f2 q when k = l + q))" by (simp add: Sum_any.expand_superset [of "{q. f2 q \ 0 \ k = l + q}"]) } note aux = this have "{k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0} \ {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto elim!: Sum_any.not_neutral_obtains_not_neutral) also have "\ \ {k. (\l. f1 l \ 0 \ sum f2 {q. f2 q \ 0 \ k = l + q} \ 0)}" by (auto dest: mult_not_zero) also have "\ \ {k. (\l. f1 l \ 0 \ (\q. f2 q \ 0 \ k = l + q))}" by (auto elim!: sum.not_neutral_contains_not_neutral) finally have "finite {k. (\l. f1 l * sum f2 {q. f2 q \ 0 \ k = l + q}) \ 0}" using * by (rule finite_subset) with aux have "finite {k. (\l. f1 l * (\q. (f2 q when k = l + q))) \ 0}" by simp with fin2 show ?thesis by (simp add: prod_fun_def) qed instantiation poly_mapping :: (monoid_add, semiring_0) semiring_0 begin lift_definition times_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ 'a \\<^sub>0 'b" is prod_fun by(rule finite_prod_fun) instance proof fix a b c :: "'a \\<^sub>0 'b" show "a * b * c = a * (b * c)" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" (is "finite ?F") assume fin_g: "finite {b. g b \ 0}" (is "finite ?G") assume fin_h: "finite {c. h c \ 0}" (is "finite ?H") from fin_f fin_g have fin_fg: "finite {(a, b). f a * g b \ 0}" (is "finite ?FG") by (rule finite_mult_not_eq_zero_prodI) from fin_g fin_h have fin_gh: "finite {(b, c). g b * h c \ 0}" (is "finite ?GH") by (rule finite_mult_not_eq_zero_prodI) from fin_f fin_g have fin_fg': "finite {a + b | a b. f a * g b \ 0}" (is "finite ?FG'") by (rule finite_mult_not_eq_zero_sumI) then have fin_fg'': "finite {d. (\(a, b). f a * g b when d = a + b) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) from fin_g fin_h have fin_gh': "finite {b + c | b c. g b * h c \ 0}" (is "finite ?GH'") by (rule finite_mult_not_eq_zero_sumI) then have fin_gh'': "finite {d. (\(b, c). g b * h c when d = b + c) \ 0}" by (auto intro: finite_Sum_any_not_eq_zero_weakenI) show "prod_fun (prod_fun f g) h = prod_fun f (prod_fun g h)" (is "?lhs = ?rhs") proof fix k from fin_f fin_g fin_h fin_fg'' have "?lhs k = (\(ab, c). (\(a, b). f a * g b when ab = a + b) * h c when k = ab + c)" by (simp add: prod_fun_unfold_prod) also have "\ = (\(ab, c). (\(a, b). f a * g b * h c when k = ab + c when ab = a + b))" - apply (subst Sum_any_left_distrib) - using fin_fg apply (simp add: split_def) - apply (subst Sum_any_when_independent [symmetric]) - apply (simp add: when_when when_mult mult_when split_def conj_commute) + using fin_fg + apply (simp add: Sum_any_left_distrib split_def flip: Sum_any_when_independent) + apply (simp add: when_when when_mult mult_when conj_commute) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = ab + c when ab = a + b)" apply (subst Sum_any.cartesian_product2 [of "(?FG' \ ?H) \ ?FG"]) - apply (auto simp add: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero) + apply (auto simp: finite_cartesian_product_iff fin_fg fin_fg' fin_h dest: mult_not_zero) done also have "\ = (\(ab, c, a, b). f a * g b * h c when k = a + b + c when ab = a + b)" by (rule Sum_any.cong) (simp add: split_def when_def) also have "\ = (\(ab, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when ab = (case cab of (c, a, b) \ a + b))" by (simp add: split_def) also have "\ = (\(c, a, b). f a * g b * h c when k = a + b + c)" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, cab). (case cab of (c, a, b) \ f a * g b * h c when k = a + b + c) when bc = (case cab of (c, a, b) \ b + c))" by (simp add: Sum_any_when_dependent_prod_left) also have "\ = (\(bc, c, a, b). f a * g b * h c when k = a + b + c when bc = b + c)" by (simp add: split_def) also have "\ = (\(bc, c, a, b). f a * g b * h c when bc = b + c when k = a + bc)" by (rule Sum_any.cong) (simp add: split_def when_def ac_simps) also have "\ = (\(a, bc, b, c). f a * g b * h c when bc = b + c when k = a + bc)" proof - have "bij (\(a, d, b, c). (d, c, a, b))" by (auto intro!: bijI injI surjI [of _ "\(d, c, a, b). (a, d, b, c)"] simp add: split_def) then show ?thesis by (rule Sum_any.reindex_cong) auto qed also have "\ = (\(a, bc). (\(b, c). f a * g b * h c when bc = b + c when k = a + bc))" apply (subst Sum_any.cartesian_product2 [of "(?F \ ?GH') \ ?GH"]) - apply (auto simp add: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero) + apply (auto simp: finite_cartesian_product_iff fin_f fin_gh fin_gh' ac_simps dest: mult_not_zero) done also have "\ = (\(a, bc). f a * (\(b, c). g b * h c when bc = b + c) when k = a + bc)" apply (subst Sum_any_right_distrib) using fin_gh apply (simp add: split_def) apply (subst Sum_any_when_independent [symmetric]) apply (simp add: when_when when_mult mult_when split_def ac_simps) done also from fin_f fin_g fin_h fin_gh'' have "\ = ?rhs k" by (simp add: prod_fun_unfold_prod) finally show "?lhs k = ?rhs k" . qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" apply (rule ext) - apply (auto simp add: prod_fun_def algebra_simps) - apply (subst Sum_any.distrib) - using fin_f fin_g apply (auto intro: finite_mult_not_eq_zero_rightI) - done + apply (simp add: prod_fun_def algebra_simps) + by (simp add: Sum_any.distrib fin_f fin_g finite_mult_not_eq_zero_rightI) qed show "a * (b + c) = a * b + a * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun f (\k. g k + h k) = (\k. prod_fun f g k + prod_fun f h k)" apply (rule ext) - apply (auto simp add: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib) - apply (subst Sum_any.distrib) - apply (simp_all add: algebra_simps) - apply (auto intro: fin_g fin_h) - apply (subst Sum_any.distrib) - apply (simp_all add: algebra_simps) - using fin_f apply (rule finite_mult_not_eq_zero_rightI) - using fin_f apply (rule finite_mult_not_eq_zero_rightI) - done + apply (auto simp: prod_fun_def Sum_any.distrib algebra_simps when_add_distrib fin_g fin_h) + by (simp add: Sum_any.distrib fin_f finite_mult_not_eq_zero_rightI) qed show "0 * a = 0" by transfer (simp add: prod_fun_def [abs_def]) show "a * 0 = 0" by transfer (simp add: prod_fun_def [abs_def]) qed end lemma lookup_mult: "lookup (f * g) k = (\l. lookup f l * (\q. lookup g q when k = l + q))" by transfer (simp add: prod_fun_def) instance poly_mapping :: (comm_monoid_add, comm_semiring_0) comm_semiring_0 proof fix a b c :: "'a \\<^sub>0 'b" show "a * b = b * a" proof transfer fix f g :: "'a \ 'b" assume fin_f: "finite {a. f a \ 0}" assume fin_g: "finite {b. g b \ 0}" show "prod_fun f g = prod_fun g f" proof fix k have fin1: "\l. finite {a. (f a when k = l + a) \ 0}" using fin_f by auto have fin2: "\l. finite {b. (g b when k = l + b) \ 0}" using fin_g by auto from fin_f fin_g have "finite {(a, b). f a \ 0 \ g b \ 0}" (is "finite ?AB") by simp - show "prod_fun f g k = prod_fun g f k" - apply (simp add: prod_fun_def) - apply (subst Sum_any_right_distrib) - apply (rule fin2) - apply (subst Sum_any_right_distrib) - apply (rule fin1) - apply (subst Sum_any.swap [of ?AB]) - apply (fact \finite ?AB\) - apply (auto simp add: mult_when ac_simps) - done + have "(\a. \n. f a * (g n when k = a + n)) = (\a. \n. g a * (f n when k = a + n))" + by (subst Sum_any.swap [OF \finite ?AB\]) (auto simp: mult_when ac_simps) + then show "prod_fun f g k = prod_fun g f k" + by (simp add: prod_fun_def Sum_any_right_distrib [OF fin2] Sum_any_right_distrib [OF fin1]) qed qed show "(a + b) * c = a * c + b * c" proof transfer fix f g h :: "'a \ 'b" assume fin_f: "finite {k. f k \ 0}" assume fin_g: "finite {k. g k \ 0}" assume fin_h: "finite {k. h k \ 0}" show "prod_fun (\k. f k + g k) h = (\k. prod_fun f h k + prod_fun g h k)" - apply (auto simp add: prod_fun_def fun_eq_iff algebra_simps) - apply (subst Sum_any.distrib) - using fin_f apply (rule finite_mult_not_eq_zero_rightI) - using fin_g apply (rule finite_mult_not_eq_zero_rightI) - apply simp_all - done + by (auto simp: prod_fun_def fun_eq_iff algebra_simps + Sum_any.distrib fin_f fin_g finite_mult_not_eq_zero_rightI) qed qed instance poly_mapping :: (monoid_add, semiring_0_cancel) semiring_0_cancel .. instance poly_mapping :: (comm_monoid_add, comm_semiring_0_cancel) comm_semiring_0_cancel .. instance poly_mapping :: (monoid_add, semiring_1) semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def] when_mult) show "a * 1 = a" apply transfer apply (simp add: prod_fun_def [abs_def] Sum_any_right_distrib Sum_any_left_distrib mult_when) apply (subst when_commute) apply simp done qed instance poly_mapping :: (comm_monoid_add, comm_semiring_1) comm_semiring_1 proof fix a :: "'a \\<^sub>0 'b" show "1 * a = a" by transfer (simp add: prod_fun_def [abs_def]) qed instance poly_mapping :: (monoid_add, semiring_1_cancel) semiring_1_cancel .. instance poly_mapping :: (monoid_add, ring) ring .. instance poly_mapping :: (comm_monoid_add, comm_ring) comm_ring .. instance poly_mapping :: (monoid_add, ring_1) ring_1 .. instance poly_mapping :: (comm_monoid_add, comm_ring_1) comm_ring_1 .. subsection \Single-point mappings\ lift_definition single :: "'a \ 'b \ 'a \\<^sub>0 'b::zero" is "\k v k'. (v when k = k')" by simp lemma inj_single [iff]: "inj (single k)" proof (rule injI, transfer) fix k :: 'b and a b :: "'a::zero" assume "(\k'. a when k = k') = (\k'. b when k = k')" then have "(\k'. a when k = k') k = (\k'. b when k = k') k" by (rule arg_cong) then show "a = b" by simp qed lemma lookup_single: "lookup (single k v) k' = (v when k = k')" - by transfer rule + by (simp add: single.rep_eq) lemma lookup_single_eq [simp]: "lookup (single k v) k = v" by transfer simp lemma lookup_single_not_eq: "k \ k' \ lookup (single k v) k' = 0" by transfer simp lemma single_zero [simp]: "single k 0 = 0" by transfer simp lemma single_one [simp]: "single 0 1 = 1" by transfer simp lemma single_add: "single k (a + b) = single k a + single k b" by transfer (simp add: fun_eq_iff when_add_distrib) lemma single_uminus: "single k (- a) = - single k a" by transfer (simp add: fun_eq_iff when_uminus_distrib) lemma single_diff: "single k (a - b) = single k a - single k b" by transfer (simp add: fun_eq_iff when_diff_distrib) lemma single_numeral [simp]: "single 0 (numeral n) = numeral n" by (induct n) (simp_all only: numeral.simps numeral_add single_zero single_one single_add) lemma lookup_numeral: "lookup (numeral n) k = (numeral n when k = 0)" proof - have "lookup (numeral n) k = lookup (single 0 (numeral n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma single_of_nat [simp]: "single 0 (of_nat n) = of_nat n" by (induct n) (simp_all add: single_add) lemma lookup_of_nat: "lookup (of_nat n) k = (of_nat n when k = 0)" proof - have "lookup (of_nat n) k = lookup (single 0 (of_nat n)) k" by simp then show ?thesis unfolding lookup_single by simp qed lemma of_nat_single: "of_nat = single 0 \ of_nat" by (simp add: fun_eq_iff) lemma mult_single: "single k a * single l b = single (k + l) (a * b)" proof transfer fix k l :: 'a and a b :: 'b show "prod_fun (\k'. a when k = k') (\k'. b when l = k') = (\k'. a * b when k + l = k')" proof fix k' have "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\n. a * b when l = n when k' = k + n)" by (simp add: prod_fun_def Sum_any_right_distrib mult_when when_mult) also have "\ = (\n. a * b when k' = k + n when l = n)" by (simp add: when_when conj_commute) also have "\ = (a * b when k' = k + l)" by simp also have "\ = (a * b when k + l = k')" by (simp add: when_def) finally show "prod_fun (\k'. a when k = k') (\k'. b when l = k') k' = (\k'. a * b when k + l = k') k'" . qed qed instance poly_mapping :: (monoid_add, semiring_char_0) semiring_char_0 by intro_classes (auto intro: inj_compose inj_of_nat simp add: of_nat_single) instance poly_mapping :: (monoid_add, ring_char_0) ring_char_0 .. lemma single_of_int [simp]: "single 0 (of_int k) = of_int k" by (cases k) (simp_all add: single_diff single_uminus) lemma lookup_of_int: "lookup (of_int l) k = (of_int l when k = 0)" -proof - - have "lookup (of_int l) k = lookup (single 0 (of_int l)) k" - by simp - then show ?thesis unfolding lookup_single by simp -qed + by (metis lookup_single_not_eq single.rep_eq single_of_int) subsection \Integral domains\ instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", semiring_no_zero_divisors) semiring_no_zero_divisors text \The @{class "linorder"} constraint is a pragmatic device for the proof --- maybe it can be dropped\ proof fix f g :: "'a \\<^sub>0 'b" assume "f \ 0" and "g \ 0" then show "f * g \ 0" proof transfer fix f g :: "'a \ 'b" define F where "F = {a. f a \ 0}" moreover define G where "G = {a. g a \ 0}" ultimately have [simp]: "\a. f a \ 0 \ a \ F" "\b. g b \ 0 \ b \ G" by simp_all assume "finite {a. f a \ 0}" then have [simp]: "finite F" by simp assume "finite {a. g a \ 0}" then have [simp]: "finite G" by simp assume "f \ (\a. 0)" then obtain a where "f a \ 0" - by (auto simp add: fun_eq_iff) + by (auto simp: fun_eq_iff) assume "g \ (\b. 0)" then obtain b where "g b \ 0" - by (auto simp add: fun_eq_iff) + by (auto simp: fun_eq_iff) from \f a \ 0\ and \g b \ 0\ have "F \ {}" and "G \ {}" by auto note Max_F = \finite F\ \F \ {}\ note Max_G = \finite G\ \G \ {}\ from Max_F and Max_G have [simp]: "Max F \ F" "Max G \ G" by auto from Max_F Max_G have [dest!]: "\a. a \ F \ a \ Max F" "\b. b \ G \ b \ Max G" by auto define q where "q = Max F + Max G" have "(\(a, b). f a * g b when q = a + b) = (\(a, b). f a * g b when q = a + b when a \ F \ b \ G)" - by (rule Sum_any.cong) (auto simp add: split_def when_def q_def intro: ccontr) + by (rule Sum_any.cong) (auto simp: split_def when_def q_def intro: ccontr) also have "\ = (\(a, b). f a * g b when (Max F, Max G) = (a, b))" proof (rule Sum_any.cong) fix ab :: "'a \ 'a" obtain a b where [simp]: "ab = (a, b)" by (cases ab) simp_all have [dest!]: "a \ Max F \ Max F \ a \ a < Max F" "b \ Max G \ Max G \ b \ b < Max G" by auto show "(case ab of (a, b) \ f a * g b when q = a + b when a \ F \ b \ G) = (case ab of (a, b) \ f a * g b when (Max F, Max G) = (a, b))" - by (auto simp add: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"]) + by (auto simp: split_def when_def q_def dest: add_strict_mono [of a "Max F" b "Max G"]) qed also have "\ = (\ab. (case ab of (a, b) \ f a * g b) when (Max F, Max G) = ab)" unfolding split_def when_def by auto also have "\ \ 0" by simp finally have "prod_fun f g q \ 0" by (simp add: prod_fun_unfold_prod) then show "prod_fun f g \ (\k. 0)" - by (auto simp add: fun_eq_iff) + by (auto simp: fun_eq_iff) qed qed instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_no_zero_divisors) ring_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", ring_1_no_zero_divisors) ring_1_no_zero_divisors .. instance poly_mapping :: ("{ordered_cancel_comm_monoid_add, linorder}", idom) idom .. subsection \Mapping order\ instantiation poly_mapping :: (linorder, "{zero, linorder}") linorder begin lift_definition less_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is less_fun . lift_definition less_eq_poly_mapping :: "('a \\<^sub>0 'b) \ ('a \\<^sub>0 'b) \ bool" is "\f g. less_fun f g \ f = g" . instance proof (rule linorder.intro_of_class) show "class.linorder (less_eq :: (_ \\<^sub>0 _) \ _) less" proof (rule linorder_strictI, rule order_strictI) fix f g h :: "'a \\<^sub>0 'b" show "f \ g \ f < g \ f = g" by transfer (rule refl) show "\ f < f" by transfer (rule less_fun_irrefl) show "f < g \ f = g \ g < f" proof transfer fix f g :: "'a \ 'b" assume "finite {k. f k \ 0}" and "finite {k. g k \ 0}" then have "finite ({k. f k \ 0} \ {k. g k \ 0})" by simp moreover have "{k. f k \ g k} \ {k. f k \ 0} \ {k. g k \ 0}" by auto ultimately have "finite {k. f k \ g k}" by (rule rev_finite_subset) then show "less_fun f g \ f = g \ less_fun g f" by (rule less_fun_trichotomy) qed assume "f < g" then show "\ g < f" by transfer (rule less_fun_asym) note \f < g\ moreover assume "g < h" ultimately show "f < h" by transfer (rule less_fun_trans) qed qed end instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, linorder}") ordered_ab_semigroup_add proof (intro_classes, transfer) fix f g h :: "'a \ 'b" assume *: "less_fun f g \ f = g" { assume "less_fun f g" then obtain k where "f k < g k" "(\k'. k' < k \ f k' = g k')" by (blast elim!: less_funE) then have "h k + f k < h k + g k" "(\k'. k' < k \ h k' + f k' = h k' + g k')" by simp_all then have "less_fun (\k. h k + f k) (\k. h k + g k)" by (blast intro: less_funI) } with * show "less_fun (\k. h k + f k) (\k. h k + g k) \ (\k. h k + f k) = (\k. h k + g k)" - by (auto simp add: fun_eq_iff) + by (auto simp: fun_eq_iff) qed instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") linordered_cancel_ab_semigroup_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_comm_monoid_add .. instance poly_mapping :: (linorder, "{ordered_comm_monoid_add, ordered_ab_semigroup_add_imp_le, cancel_comm_monoid_add, linorder}") ordered_cancel_comm_monoid_add .. instance poly_mapping :: (linorder, linordered_ab_group_add) linordered_ab_group_add .. text \ For pragmatism we leave out the final elements in the hierarchy: @{class linordered_ring}, @{class linordered_ring_strict}, @{class linordered_idom}; remember that the order instance is a mere technical device, not a deeper algebraic property. \ subsection \Fundamental mapping notions\ lift_definition keys :: "('a \\<^sub>0 'b::zero) \ 'a set" is "\f. {k. f k \ 0}" . lift_definition range :: "('a \\<^sub>0 'b::zero) \ 'b set" is "\f :: 'a \ 'b. Set.range f - {0}" . lemma finite_keys [simp]: "finite (keys f)" by transfer lemma not_in_keys_iff_lookup_eq_zero: "k \ keys f \ lookup f k = 0" by transfer simp lemma lookup_not_eq_zero_eq_in_keys: "lookup f k \ 0 \ k \ keys f" by transfer simp lemma lookup_eq_zero_in_keys_contradict [dest]: "lookup f k = 0 \ \ k \ keys f" by (simp add: not_in_keys_iff_lookup_eq_zero) lemma finite_range [simp]: "finite (Poly_Mapping.range p)" proof transfer fix f :: "'b \ 'a" assume *: "finite {x. f x \ 0}" have "Set.range f - {0} \ f ` {x. f x \ 0}" by auto thus "finite (Set.range f - {0})" by(rule finite_subset)(rule finite_imageI[OF *]) qed lemma in_keys_lookup_in_range [simp]: "k \ keys f \ lookup f k \ range f" by transfer simp lemma in_keys_iff: "x \ (keys s) = (lookup s x \ 0)" by (transfer, simp) lemma keys_zero [simp]: "keys 0 = {}" by transfer simp lemma range_zero [simp]: "range 0 = {}" by transfer auto lemma keys_add: "keys (f + g) \ keys f \ keys g" by transfer auto lemma keys_one [simp]: "keys 1 = {0}" by transfer simp lemma range_one [simp]: "range 1 = {1}" - by transfer (auto simp add: when_def) + by transfer (auto simp: when_def) lemma keys_single [simp]: "keys (single k v) = (if v = 0 then {} else {k})" by transfer simp lemma range_single [simp]: "range (single k v) = (if v = 0 then {} else {v})" - by transfer (auto simp add: when_def) + by transfer (auto simp: when_def) lemma keys_mult: "keys (f * g) \ {a + b | a b. a \ keys f \ b \ keys g}" apply transfer - apply (auto simp add: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral) - apply blast + apply (force simp: prod_fun_def dest!: mult_not_zero elim!: Sum_any.not_neutral_obtains_not_neutral) done lemma setsum_keys_plus_distrib: assumes hom_0: "\k. f k 0 = 0" and hom_plus: "\k. k \ Poly_Mapping.keys p \ Poly_Mapping.keys q \ f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k) = f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k)" shows "(\k\Poly_Mapping.keys (p + q). f k (Poly_Mapping.lookup (p + q) k)) = (\k\Poly_Mapping.keys p. f k (Poly_Mapping.lookup p k)) + (\k\Poly_Mapping.keys q. f k (Poly_Mapping.lookup q k))" (is "?lhs = ?p + ?q") proof - let ?A = "Poly_Mapping.keys p \ Poly_Mapping.keys q" have "?lhs = (\k\?A. f k (Poly_Mapping.lookup p k + Poly_Mapping.lookup q k))" - apply(rule sum.mono_neutral_cong_left) - apply(simp_all add: Poly_Mapping.keys_add) - apply(transfer fixing: f) - apply(auto simp add: hom_0)[1] - apply(transfer fixing: f) - apply(auto simp add: hom_0)[1] - done + by(intro sum.mono_neutral_cong_left) (auto simp: sum.mono_neutral_cong_left hom_0 in_keys_iff lookup_add) also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k) + f k (Poly_Mapping.lookup q k))" by(rule sum.cong)(simp_all add: hom_plus) also have "\ = (\k\?A. f k (Poly_Mapping.lookup p k)) + (\k\?A. f k (Poly_Mapping.lookup q k))" (is "_ = ?p' + ?q'") by(simp add: sum.distrib) also have "?p' = ?p" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) also have "?q' = ?q" by (simp add: hom_0 in_keys_iff sum.mono_neutral_cong_right) finally show ?thesis . qed subsection \Degree\ definition degree :: "(nat \\<^sub>0 'a::zero) \ nat" where "degree f = Max (insert 0 (Suc ` keys f))" lemma degree_zero [simp]: "degree 0 = 0" unfolding degree_def by transfer simp lemma degree_one [simp]: "degree 1 = 1" unfolding degree_def by transfer simp lemma degree_single_zero [simp]: "degree (single k 0) = 0" unfolding degree_def by transfer simp lemma degree_single_not_zero [simp]: "v \ 0 \ degree (single k v) = Suc k" unfolding degree_def by transfer simp lemma degree_zero_iff [simp]: "degree f = 0 \ f = 0" unfolding degree_def proof transfer fix f :: "nat \ 'a" assume "finite {n. f n \ 0}" then have fin: "finite (insert 0 (Suc ` {n. f n \ 0}))" by auto show "Max (insert 0 (Suc ` {n. f n \ 0})) = 0 \ f = (\n. 0)" (is "?P \ ?Q") proof assume ?P have "{n. f n \ 0} = {}" proof (rule ccontr) assume "{n. f n \ 0} \ {}" then obtain n where "n \ {n. f n \ 0}" by blast then have "{n. f n \ 0} = insert n {n. f n \ 0}" by auto then have "Suc ` {n. f n \ 0} = insert (Suc n) (Suc ` {n. f n \ 0})" by auto with \?P\ have "Max (insert 0 (insert (Suc n) (Suc ` {n. f n \ 0}))) = 0" by simp then have "Max (insert (Suc n) (insert 0 (Suc ` {n. f n \ 0}))) = 0" by (simp add: insert_commute) with fin have "max (Suc n) (Max (insert 0 (Suc ` {n. f n \ 0}))) = 0" by simp then show False by simp qed then show ?Q by (simp add: fun_eq_iff) next assume ?Q then show ?P by simp qed qed lemma degree_greater_zero_in_keys: assumes "0 < degree f" shows "degree f - 1 \ keys f" proof - from assms have "keys f \ {}" - by (auto simp add: degree_def) + by (auto simp: degree_def) then show ?thesis unfolding degree_def by (simp add: mono_Max_commute [symmetric] mono_Suc) qed lemma in_keys_less_degree: "n \ keys f \ n < degree f" -unfolding degree_def by transfer (auto simp add: Max_gr_iff) +unfolding degree_def by transfer (auto simp: Max_gr_iff) lemma beyond_degree_lookup_zero: "degree f \ n \ lookup f n = 0" unfolding degree_def by transfer auto lemma degree_add: "degree (f + g) \ max (degree f) (Poly_Mapping.degree g)" unfolding degree_def proof transfer fix f g :: "nat \ 'a" assume f: "finite {x. f x \ 0}" assume g: "finite {x. g x \ 0}" let ?f = "Max (insert 0 (Suc ` {k. f k \ 0}))" let ?g = "Max (insert 0 (Suc ` {k. g k \ 0}))" have "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ Max (insert 0 (Suc ` ({k. f k \ 0} \ {k. g k \ 0})))" by (rule Max.subset_imp) (insert f g, auto) also have "\ = max ?f ?g" using f g by (simp_all add: image_Un Max_Un [symmetric]) finally show "Max (insert 0 (Suc ` {k. f k + g k \ 0})) \ max (Max (insert 0 (Suc ` {k. f k \ 0}))) (Max (insert 0 (Suc ` {k. g k \ 0})))" . qed lemma sorted_list_of_set_keys: "sorted_list_of_set (keys f) = filter (\k. k \ keys f) [0..Inductive structure\ lift_definition update :: "'a \ 'b \ ('a \\<^sub>0 'b::zero) \ 'a \\<^sub>0 'b" is "\k v f. f(k := v)" proof - fix f :: "'a \ 'b" and k' v assume "finite {k. f k \ 0}" then have "finite (insert k' {k. f k \ 0})" by simp then show "finite {k. (f(k' := v)) k \ 0}" by (rule rev_finite_subset) auto qed lemma update_induct [case_names const update]: assumes const': "P 0" assumes update': "\f a b. a \ keys f \ b \ 0 \ P f \ P (update a b f)" shows "P f" proof - obtain g where "f = Abs_poly_mapping g" and "finite {a. g a \ 0}" by (cases f) simp_all define Q where "Q g = P (Abs_poly_mapping g)" for g from \finite {a. g a \ 0}\ have "Q g" proof (induct g rule: finite_update_induct) case const with const' Q_def show ?case by simp next case (update a b g) from \finite {a. g a \ 0}\ \g a = 0\ have "a \ keys (Abs_poly_mapping g)" by (simp add: Abs_poly_mapping_inverse keys.rep_eq) moreover note \b \ 0\ moreover from \Q g\ have "P (Abs_poly_mapping g)" by (simp add: Q_def) ultimately have "P (update a b (Abs_poly_mapping g))" by (rule update') also from \finite {a. g a \ 0}\ have "update a b (Abs_poly_mapping g) = Abs_poly_mapping (g(a := b))" by (simp add: update.abs_eq eq_onp_same_args) finally show ?case by (simp add: Q_def fun_upd_def) qed then show ?thesis by (simp add: Q_def \f = Abs_poly_mapping g\) qed lemma lookup_update: "lookup (update k v f) k' = (if k = k' then v else lookup f k')" by transfer simp lemma keys_update: "keys (update k v f) = (if v = 0 then keys f - {k} else insert k (keys f))" by transfer auto subsection \Quasi-functorial structure\ lift_definition map :: "('b::zero \ 'c::zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c::zero)" is "\g f k. g (f k) when f k \ 0" by simp context fixes f :: "'b \ 'a" assumes inj_f: "inj f" begin lift_definition map_key :: "('a \\<^sub>0 'c::zero) \ 'b \\<^sub>0 'c" is "\p. p \ f" proof - fix g :: "'c \ 'd" and p :: "'a \ 'c" assume "finite {x. p x \ 0}" hence "finite (f ` {y. p (f y) \ 0})" by(rule finite_subset[rotated]) auto thus "finite {x. (p \ f) x \ 0}" unfolding o_def by(rule finite_imageD)(rule subset_inj_on[OF inj_f], simp) qed end lemma map_key_compose: assumes [transfer_rule]: "inj f" "inj g" shows "map_key f (map_key g p) = map_key (g \ f) p" proof - from assms have [transfer_rule]: "inj (g \ f)" by(simp add: inj_compose) show ?thesis by transfer(simp add: o_assoc) qed lemma map_key_id: "map_key (\x. x) p = p" proof - have [transfer_rule]: "inj (\x. x)" by simp show ?thesis by transfer(simp add: o_def) qed context fixes f :: "'a \ 'b" assumes inj_f [transfer_rule]: "inj f" begin lemma map_key_map: "map_key f (map g p) = map g (map_key f p)" by transfer (simp add: fun_eq_iff) lemma map_key_plus: "map_key f (p + q) = map_key f p + map_key f q" by transfer (simp add: fun_eq_iff) lemma keys_map_key: "keys (map_key f p) = f -` keys p" by transfer auto lemma map_key_zero [simp]: "map_key f 0 = 0" by transfer (simp add: fun_eq_iff) lemma map_key_single [simp]: "map_key f (single (f k) v) = single k v" by transfer (simp add: fun_eq_iff inj_onD [OF inj_f] when_def) end lemma mult_map_scale_conv_mult: "map ((*) s) p = single 0 s * p" proof(transfer fixing: s) fix p :: "'a \ 'b" assume *: "finite {x. p x \ 0}" { fix x have "prod_fun (\k'. s when 0 = k') p x = (\l :: 'a. if l = 0 then s * (\q. p q when x = q) else 0)" - by(auto simp add: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta) + by(auto simp: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta) also have "\ = (\k. s * p k when p k \ 0) x" by(simp add: when_def) also note calculation } then show "(\k. s * p k when p k \ 0) = prod_fun (\k'. s when 0 = k') p" by(simp add: fun_eq_iff) qed lemma map_single [simp]: "(c = 0 \ f 0 = 0) \ map f (single x c) = single x (f c)" -by transfer(auto simp add: fun_eq_iff when_def) + by transfer(auto simp: fun_eq_iff when_def) lemma map_eq_zero_iff: "map f p = 0 \ (\k \ keys p. f (lookup p k) = 0)" -by transfer(auto simp add: fun_eq_iff when_def) + by transfer(auto simp: fun_eq_iff when_def) subsection \Canonical dense representation of @{typ "nat \\<^sub>0 'a"}\ abbreviation no_trailing_zeros :: "'a :: zero list \ bool" where "no_trailing_zeros \ no_trailing ((=) 0)" lift_definition "nth" :: "'a list \ (nat \\<^sub>0 'a::zero)" is "nth_default 0" by (fact finite_nth_default_neq_default) text \ The opposite direction is directly specified on (later) type \nat_mapping\. \ lemma nth_Nil [simp]: "nth [] = 0" by transfer (simp add: fun_eq_iff) lemma nth_singleton [simp]: "nth [v] = single 0 v" proof (transfer, rule ext) fix n :: nat and v :: 'a show "nth_default 0 [v] n = (v when 0 = n)" - by (auto simp add: nth_default_def nth_append) + by (auto simp: nth_default_def nth_append) qed lemma nth_replicate [simp]: "nth (replicate n 0 @ [v]) = single n v" proof (transfer, rule ext) fix m n :: nat and v :: 'a show "nth_default 0 (replicate n 0 @ [v]) m = (v when n = m)" - by (auto simp add: nth_default_def nth_append) + by (auto simp: nth_default_def nth_append) qed lemma nth_strip_while [simp]: "nth (strip_while ((=) 0) xs) = nth xs" by transfer (fact nth_default_strip_while_dflt) lemma nth_strip_while' [simp]: "nth (strip_while (\k. k = 0) xs) = nth xs" by (subst eq_commute) (fact nth_strip_while) lemma nth_eq_iff: "nth xs = nth ys \ strip_while (HOL.eq 0) xs = strip_while (HOL.eq 0) ys" by transfer (simp add: nth_default_eq_iff) lemma lookup_nth [simp]: "lookup (nth xs) = nth_default 0 xs" by (fact nth.rep_eq) lemma keys_nth [simp]: "keys (nth xs) = fst ` {(n, v) \ set (enumerate 0 xs). v \ 0}" proof transfer fix xs :: "'a list" { fix n assume "nth_default 0 xs n \ 0" then have "n < length xs" and "xs ! n \ 0" - by (auto simp add: nth_default_def split: if_splits) + by (auto simp: nth_default_def split: if_splits) then have "(n, xs ! n) \ {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" (is "?x \ ?A") - by (auto simp add: in_set_conv_nth enumerate_eq_zip) + by (auto simp: in_set_conv_nth enumerate_eq_zip) then have "fst ?x \ fst ` ?A" by blast then have "n \ fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" by simp } then show "{k. nth_default 0 xs k \ 0} = fst ` {(n, v). (n, v) \ set (enumerate 0 xs) \ v \ 0}" - by (auto simp add: in_enumerate_iff_nth_default_eq) + by (auto simp: in_enumerate_iff_nth_default_eq) qed lemma range_nth [simp]: "range (nth xs) = set xs - {0}" by transfer simp lemma degree_nth: "no_trailing_zeros xs \ degree (nth xs) = length xs" unfolding degree_def proof transfer fix xs :: "'a list" assume *: "no_trailing_zeros xs" let ?A = "{n. nth_default 0 xs n \ 0}" let ?f = "nth_default 0 xs" let ?bound = "Max (insert 0 (Suc ` {n. ?f n \ 0}))" show "?bound = length xs" proof (cases "xs = []") case False with * obtain n where n: "n < length xs" "xs ! n \ 0" by (fastforce simp add: no_trailing_unfold last_conv_nth neq_Nil_conv) then have "?bound = Max (Suc ` {k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs})" - by (subst Max_insert) (auto simp add: nth_default_def) + by (subst Max_insert) (auto simp: nth_default_def) also let ?A = "{k. k < length xs \ xs ! k \ 0}" have "{k. (k < length xs \ xs ! k \ (0::'a)) \ k < length xs} = ?A" by auto also have "Max (Suc ` ?A) = Suc (Max ?A)" using n - by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp add: mono_Suc) + by (subst mono_Max_commute [where f = Suc, symmetric]) (auto simp: mono_Suc) also { have "Max ?A \ ?A" using n Max_in [of ?A] by fastforce hence "Suc (Max ?A) \ length xs" by simp moreover from * False have "length xs - 1 \ ?A" - by(auto simp add: no_trailing_unfold last_conv_nth) + by(auto simp: no_trailing_unfold last_conv_nth) hence "length xs - 1 \ Max ?A" using Max_ge[of ?A "length xs - 1"] by auto hence "length xs \ Suc (Max ?A)" by simp ultimately have "Suc (Max ?A) = length xs" by simp } finally show ?thesis . qed simp qed lemma nth_trailing_zeros [simp]: "nth (xs @ replicate n 0) = nth xs" by transfer simp lemma nth_idem: "nth (List.map (lookup f) [0.. n" shows "nth (List.map (lookup f) [0..Canonical sparse representation of @{typ "'a \\<^sub>0 'b"}\ lift_definition the_value :: "('a \ 'b) list \ 'a \\<^sub>0 'b::zero" is "\xs k. case map_of xs k of None \ 0 | Some v \ v" proof - fix xs :: "('a \ 'b) list" have fin: "finite {k. \v. map_of xs k = Some v}" using finite_dom_map_of [of xs] unfolding dom_def by auto then show "finite {k. (case map_of xs k of None \ 0 | Some v \ v) \ 0}" using fin by (simp split: option.split) qed definition items :: "('a::linorder \\<^sub>0 'b::zero) \ ('a \ 'b) list" where "items f = List.map (\k. (k, lookup f k)) (sorted_list_of_set (keys f))" text \ For the canonical sparse representation we provide both directions of morphisms since the specification of ordered association lists in theory \OAList\ will support arbitrary linear orders @{class linorder} as keys, not just natural numbers @{typ nat}. \ lemma the_value_items [simp]: "the_value (items f) = f" unfolding items_def by transfer (simp add: fun_eq_iff map_of_map_restrict restrict_map_def) lemma lookup_the_value: "lookup (the_value xs) k = (case map_of xs k of None \ 0 | Some v \ v)" by transfer rule lemma items_the_value: assumes "sorted (List.map fst xs)" and "distinct (List.map fst xs)" and "0 \ snd ` set xs" shows "items (the_value xs) = xs" proof - from assms have "sorted_list_of_set (set (List.map fst xs)) = List.map fst xs" unfolding sorted_list_of_set_sort_remdups by (simp add: distinct_remdups_id sort_key_id_if_sorted) moreover from assms have "keys (the_value xs) = fst ` set xs" - by transfer (auto simp add: image_def split: option.split dest: set_map_of_compr) + by transfer (auto simp: image_def split: option.split dest: set_map_of_compr) ultimately show ?thesis unfolding items_def using assms - by (auto simp add: lookup_the_value intro: map_idI) + by (auto simp: lookup_the_value intro: map_idI) qed lemma the_value_Nil [simp]: "the_value [] = 0" by transfer (simp add: fun_eq_iff) lemma the_value_Cons [simp]: "the_value (x # xs) = update (fst x) (snd x) (the_value xs)" by transfer (simp add: fun_eq_iff) lemma items_zero [simp]: "items 0 = []" unfolding items_def by simp lemma items_one [simp]: "items 1 = [(0, 1)]" unfolding items_def by transfer simp lemma items_single [simp]: "items (single k v) = (if v = 0 then [] else [(k, v)])" unfolding items_def by simp lemma in_set_items_iff [simp]: "(k, v) \ set (items f) \ k \ keys f \ lookup f k = v" unfolding items_def by transfer auto subsection \Size estimation\ context fixes f :: "'a \ nat" and g :: "'b :: zero \ nat" begin definition poly_mapping_size :: "('a \\<^sub>0 'b) \ nat" where "poly_mapping_size m = g 0 + (\k \ keys m. Suc (f k + g (lookup m k)))" lemma poly_mapping_size_0 [simp]: "poly_mapping_size 0 = g 0" by (simp add: poly_mapping_size_def) lemma poly_mapping_size_single [simp]: "poly_mapping_size (single k v) = (if v = 0 then g 0 else g 0 + f k + g v + 1)" unfolding poly_mapping_size_def by transfer simp lemma keys_less_poly_mapping_size: "k \ keys m \ f k + g (lookup m k) < poly_mapping_size m" unfolding poly_mapping_size_def proof transfer fix k :: 'a and m :: "'a \ 'b" and f :: "'a \ nat" and g let ?keys = "{k. m k \ 0}" assume *: "finite ?keys" "k \ ?keys" then have "f k + g (m k) = (\k' \ ?keys. f k' + g (m k') when k' = k)" by (simp add: sum.delta when_def) also have "\ < (\k' \ ?keys. Suc (f k' + g (m k')))" using * - by (intro sum_strict_mono) (auto simp add: when_def) + by (intro sum_strict_mono) (auto simp: when_def) also have "\ \ g 0 + \" by simp finally have "f k + g (m k) < \" . then show "f k + g (m k) < g 0 + (\k | m k \ 0. Suc (f k + g (m k)))" by simp qed lemma lookup_le_poly_mapping_size: "g (lookup m k) \ poly_mapping_size m" proof (cases "k \ keys m") case True with keys_less_poly_mapping_size [of k m] show ?thesis by simp next case False then show ?thesis by (simp add: Poly_Mapping.poly_mapping_size_def in_keys_iff) qed lemma poly_mapping_size_estimation: "k \ keys m \ y \ f k + g (lookup m k) \ y < poly_mapping_size m" using keys_less_poly_mapping_size by (auto intro: le_less_trans) lemma poly_mapping_size_estimation2: assumes "v \ range m" and "y \ g v" shows "y < poly_mapping_size m" proof - from assms obtain k where *: "lookup m k = v" "v \ 0" by transfer blast from * have "k \ keys m" by (simp add: in_keys_iff) then show ?thesis proof (rule poly_mapping_size_estimation) from assms * have "y \ g (lookup m k)" by simp then show "y \ f k + g (lookup m k)" by simp qed qed end lemma poly_mapping_size_one [simp]: "poly_mapping_size f g 1 = g 0 + f 0 + g 1 + 1" unfolding poly_mapping_size_def by transfer simp lemma poly_mapping_size_cong [fundef_cong]: "m = m' \ g 0 = g' 0 \ (\k. k \ keys m' \ f k = f' k) \ (\v. v \ range m' \ g v = g' v) \ poly_mapping_size f g m = poly_mapping_size f' g' m'" - by (auto simp add: poly_mapping_size_def intro!: sum.cong) + by (auto simp: poly_mapping_size_def intro!: sum.cong) instantiation poly_mapping :: (type, zero) size begin definition "size = poly_mapping_size (\_. 0) (\_. 0)" instance .. end subsection \Further mapping operations and properties\ text \It is like in algebra: there are many definitions, some are also used\ lift_definition mapp :: "('a \ 'b :: zero \ 'c :: zero) \ ('a \\<^sub>0 'b) \ ('a \\<^sub>0 'c)" is "\f p k. (if k \ keys p then f k (lookup p k) else 0)" by simp lemma mapp_cong [fundef_cong]: "\ m = m'; \k. k \ keys m' \ f k (lookup m' k) = f' k (lookup m' k) \ \ mapp f m = mapp f' m'" - by transfer (auto simp add: fun_eq_iff) + by transfer (auto simp: fun_eq_iff) lemma lookup_mapp: "lookup (mapp f p) k = (f k (lookup p k) when k \ keys p)" by (simp add: mapp.rep_eq) lemma keys_mapp_subset: "keys (mapp f p) \ keys p" by (meson in_keys_iff mapp.rep_eq subsetI) subsection\Free Abelian Groups Over a Type\ abbreviation frag_of :: "'a \ 'a \\<^sub>0 int" where "frag_of c \ Poly_Mapping.single c (1::int)" lemma lookup_frag_of [simp]: "Poly_Mapping.lookup(frag_of c) = (\x. if x = c then 1 else 0)" by (force simp add: lookup_single_not_eq) lemma frag_of_nonzero [simp]: "frag_of a \ 0" proof - let ?f = "\x. if x = a then 1 else (0::int)" have "?f \ (\x. 0::int)" by (auto simp: fun_eq_iff) then have "Poly_Mapping.lookup (Abs_poly_mapping ?f) \ Poly_Mapping.lookup (Abs_poly_mapping (\x. 0))" by fastforce then show ?thesis by (metis lookup_single_eq lookup_zero) qed definition frag_cmul :: "int \ ('a \\<^sub>0 int) \ ('a \\<^sub>0 int)" where "frag_cmul c a = Abs_poly_mapping (\x. c * Poly_Mapping.lookup a x)" lemma frag_cmul_zero [simp]: "frag_cmul 0 x = 0" by (simp add: frag_cmul_def) lemma frag_cmul_zero2 [simp]: "frag_cmul c 0 = 0" by (simp add: frag_cmul_def) lemma frag_cmul_one [simp]: "frag_cmul 1 x = x" by (auto simp: frag_cmul_def Poly_Mapping.poly_mapping.lookup_inverse) lemma frag_cmul_minus_one [simp]: "frag_cmul (-1) x = -x" by (simp add: frag_cmul_def uminus_poly_mapping_def poly_mapping_eqI) lemma frag_cmul_cmul [simp]: "frag_cmul c (frag_cmul d x) = frag_cmul (c*d) x" by (simp add: frag_cmul_def mult_ac) lemma lookup_frag_cmul [simp]: "poly_mapping.lookup (frag_cmul c x) i = c * poly_mapping.lookup x i" by (simp add: frag_cmul_def) lemma minus_frag_cmul [simp]: "- frag_cmul k x = frag_cmul (-k) x" by (simp add: poly_mapping_eqI) lemma keys_frag_of: "Poly_Mapping.keys(frag_of a) = {a}" by simp lemma finite_cmul_nonzero: "finite {x. c * Poly_Mapping.lookup a x \ (0::int)}" by simp lemma keys_cmul: "Poly_Mapping.keys(frag_cmul c a) \ Poly_Mapping.keys a" using finite_cmul_nonzero [of c a] by (metis lookup_frag_cmul mult_zero_right not_in_keys_iff_lookup_eq_zero subsetI) lemma keys_cmul_iff [iff]: "i \ Poly_Mapping.keys (frag_cmul c x) \ i \ Poly_Mapping.keys x \ c \ 0" - apply auto - apply (meson subsetD keys_cmul) by (metis in_keys_iff lookup_frag_cmul mult_eq_0_iff) lemma keys_minus [simp]: "Poly_Mapping.keys(-a) = Poly_Mapping.keys a" by (metis (no_types, opaque_lifting) in_keys_iff lookup_uminus neg_equal_0_iff_equal subsetI subset_antisym) lemma keys_diff: "Poly_Mapping.keys(a - b) \ Poly_Mapping.keys a \ Poly_Mapping.keys b" - by (auto simp add: in_keys_iff lookup_minus) + by (auto simp: in_keys_iff lookup_minus) lemma keys_eq_empty [simp]: "Poly_Mapping.keys c = {} \ c = 0" by (metis in_keys_iff keys_zero lookup_zero poly_mapping_eqI) lemma frag_cmul_eq_0_iff [simp]: "frag_cmul k c = 0 \ k=0 \ c=0" by auto (metis subsetI subset_antisym keys_cmul_iff keys_eq_empty) lemma frag_of_eq: "frag_of x = frag_of y \ x = y" by (metis lookup_single_eq lookup_single_not_eq zero_neq_one) lemma frag_cmul_distrib: "frag_cmul (c+d) a = frag_cmul c a + frag_cmul d a" by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) lemma frag_cmul_distrib2: "frag_cmul c (a+b) = frag_cmul c a + frag_cmul c b" proof - have "finite {x. poly_mapping.lookup a x + poly_mapping.lookup b x \ 0}" using keys_add [of a b] by (metis (no_types, lifting) finite_keys finite_subset keys.rep_eq lookup_add mem_Collect_eq subsetI) then show ?thesis by (simp add: frag_cmul_def plus_poly_mapping_def int_distrib) qed lemma frag_cmul_diff_distrib: "frag_cmul (a - b) c = frag_cmul a c - frag_cmul b c" by (auto simp: left_diff_distrib lookup_minus poly_mapping_eqI) lemma frag_cmul_sum: "frag_cmul a (sum b I) = (\i\I. frag_cmul a (b i))" proof (induction rule: infinite_finite_induct) case (insert i I) then show ?case by (auto simp: algebra_simps frag_cmul_distrib2) qed auto lemma keys_sum: "Poly_Mapping.keys(sum b I) \ (\i \I. Poly_Mapping.keys(b i))" proof (induction I rule: infinite_finite_induct) case (insert i I) then show ?case using keys_add [of "b i" "sum b I"] by auto qed auto definition frag_extend :: "('b \ 'a \\<^sub>0 int) \ ('b \\<^sub>0 int) \ 'a \\<^sub>0 int" where "frag_extend b x \ (\i \ Poly_Mapping.keys x. frag_cmul (Poly_Mapping.lookup x i) (b i))" lemma frag_extend_0 [simp]: "frag_extend b 0 = 0" by (simp add: frag_extend_def) lemma frag_extend_of [simp]: "frag_extend f (frag_of a) = f a" by (simp add: frag_extend_def) lemma frag_extend_cmul: "frag_extend f (frag_cmul c x) = frag_cmul c (frag_extend f x)" by (auto simp: frag_extend_def frag_cmul_sum intro: sum.mono_neutral_cong_left) lemma frag_extend_minus: "frag_extend f (- x) = - (frag_extend f x)" using frag_extend_cmul [of f "-1"] by simp lemma frag_extend_add: "frag_extend f (a+b) = (frag_extend f a) + (frag_extend f b)" proof - have *: "(\i\Poly_Mapping.keys a. frag_cmul (poly_mapping.lookup a i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i))" "(\i\Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i)) = (\i\Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))" by (auto simp: in_keys_iff intro: sum.mono_neutral_cong_left) have "frag_extend f (a+b) = (\i\Poly_Mapping.keys (a + b). frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i)) " by (auto simp: frag_extend_def Poly_Mapping.lookup_add frag_cmul_distrib) also have "... = (\i \ Poly_Mapping.keys a \ Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i))" - apply (rule sum.mono_neutral_cong_left) - using keys_add [of a b] - apply (auto simp add: in_keys_iff plus_poly_mapping.rep_eq frag_cmul_distrib [symmetric]) - done + proof (rule sum.mono_neutral_cong_left) + show "\i\keys a \ keys b - keys (a + b). + frag_cmul (lookup a i) (f i) + frag_cmul (lookup b i) (f i) = 0" + by (metis DiffD2 frag_cmul_distrib frag_cmul_zero in_keys_iff lookup_add) + qed (auto simp: keys_add) also have "... = (frag_extend f a) + (frag_extend f b)" by (auto simp: * sum.distrib frag_extend_def) finally show ?thesis . qed lemma frag_extend_diff: "frag_extend f (a-b) = (frag_extend f a) - (frag_extend f b)" by (metis (no_types, opaque_lifting) add_uminus_conv_diff frag_extend_add frag_extend_minus) lemma frag_extend_sum: "finite I \ frag_extend f (\i\I. g i) = sum (frag_extend f o g) I" by (induction I rule: finite_induct) (simp_all add: frag_extend_add) lemma frag_extend_eq: "(\f. f \ Poly_Mapping.keys c \ g f = h f) \ frag_extend g c = frag_extend h c" by (simp add: frag_extend_def) lemma frag_extend_eq_0: "(\x. x \ Poly_Mapping.keys c \ f x = 0) \ frag_extend f c = 0" by (simp add: frag_extend_def) lemma keys_frag_extend: "Poly_Mapping.keys(frag_extend f c) \ (\x \ Poly_Mapping.keys c. Poly_Mapping.keys(f x))" unfolding frag_extend_def using keys_sum by fastforce lemma frag_expansion: "a = frag_extend frag_of a" proof - have *: "finite I \ Poly_Mapping.lookup (\i\I. frag_cmul (Poly_Mapping.lookup a i) (frag_of i)) j = (if j \ I then Poly_Mapping.lookup a j else 0)" for I j by (induction I rule: finite_induct) (auto simp: lookup_single lookup_add) show ?thesis unfolding frag_extend_def by (rule poly_mapping_eqI) (fastforce simp add: in_keys_iff *) qed lemma frag_closure_minus_cmul: assumes "P 0" and P: "\x y. \P x; P y\ \ P(x - y)" "P c" shows "P(frag_cmul k c)" proof - have "P (frag_cmul (int n) c)" for n - apply (induction n) - apply (simp_all add: assms frag_cmul_distrib) - by (metis add.left_neutral add_diff_cancel_right' add_uminus_conv_diff P) + proof (induction n) + case 0 + then show ?case + by (simp add: assms) + next + case (Suc n) + then show ?case + by (metis assms diff_0 diff_minus_eq_add frag_cmul_distrib frag_cmul_one of_nat_Suc) + qed then show ?thesis by (metis (no_types, opaque_lifting) add_diff_eq assms(2) diff_add_cancel frag_cmul_distrib int_diff_cases) qed lemma frag_induction [consumes 1, case_names zero one diff]: assumes supp: "Poly_Mapping.keys c \ S" and 0: "P 0" and sing: "\x. x \ S \ P(frag_of x)" and diff: "\a b. \P a; P b\ \ P(a - b)" shows "P c" proof - have "P (\i\I. frag_cmul (poly_mapping.lookup c i) (frag_of i))" if "I \ Poly_Mapping.keys c" for I using finite_subset [OF that finite_keys [of c]] that supp proof (induction I arbitrary: c rule: finite_induct) case empty then show ?case by (auto simp: 0) next case (insert i I c) have ab: "a+b = a - (0 - b)" for a b :: "'a \\<^sub>0 int" by simp have Pfrag: "P (frag_cmul (poly_mapping.lookup c i) (frag_of i))" by (metis "0" diff frag_closure_minus_cmul insert.prems insert_subset sing subset_iff) - show ?case - apply (simp add: insert.hyps) - apply (subst ab) - using insert apply (blast intro: assms Pfrag) - done + with insert show ?case + by (metis (mono_tags, lifting) "0" ab diff insert_subset sum.insert) qed then show ?thesis - by (subst frag_expansion) (auto simp add: frag_extend_def) + by (subst frag_expansion) (auto simp: frag_extend_def) qed lemma frag_extend_compose: "frag_extend f (frag_extend (frag_of o g) c) = frag_extend (f o g) c" using subset_UNIV by (induction c rule: frag_induction) (auto simp: frag_extend_diff) lemma frag_split: fixes c :: "'a \\<^sub>0 int" assumes "Poly_Mapping.keys c \ S \ T" obtains d e where "Poly_Mapping.keys d \ S" "Poly_Mapping.keys e \ T" "d + e = c" proof let ?d = "frag_extend (\f. if f \ S then frag_of f else 0) c" let ?e = "frag_extend (\f. if f \ S then 0 else frag_of f) c" show "Poly_Mapping.keys ?d \ S" "Poly_Mapping.keys ?e \ T" using assms by (auto intro!: order_trans [OF keys_frag_extend] split: if_split_asm) show "?d + ?e = c" using assms proof (induction c rule: frag_induction) case (diff a b) then show ?case by (metis (no_types, lifting) frag_extend_diff add_diff_eq diff_add_eq diff_add_eq_diff_diff_swap) qed auto qed hide_const (open) lookup single update keys range map map_key degree nth the_value items foldr mapp end \ No newline at end of file