diff --git a/src/HOL/Algebra/Divisibility.thy b/src/HOL/Algebra/Divisibility.thy --- a/src/HOL/Algebra/Divisibility.thy +++ b/src/HOL/Algebra/Divisibility.thy @@ -1,2961 +1,2939 @@ (* Title: HOL/Algebra/Divisibility.thy Author: Clemens Ballarin Author: Stephan Hohe *) section \Divisibility in monoids and rings\ theory Divisibility imports "HOL-Combinatorics.List_Permutation" Coset Group begin section \Factorial Monoids\ subsection \Monoids with Cancellation Law\ locale monoid_cancel = monoid + assumes l_cancel: "\c \ a = c \ b; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" and r_cancel: "\a \ c = b \ c; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" lemma (in monoid) monoid_cancelI: assumes l_cancel: "\a b c. \c \ a = c \ b; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" and r_cancel: "\a b c. \a \ c = b \ c; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" shows "monoid_cancel G" by standard fact+ lemma (in monoid_cancel) is_monoid_cancel: "monoid_cancel G" .. sublocale group \ monoid_cancel by standard simp_all locale comm_monoid_cancel = monoid_cancel + comm_monoid lemma comm_monoid_cancelI: fixes G (structure) assumes "comm_monoid G" assumes cancel: "\a b c. \a \ c = b \ c; a \ carrier G; b \ carrier G; c \ carrier G\ \ a = b" shows "comm_monoid_cancel G" proof - interpret comm_monoid G by fact show "comm_monoid_cancel G" by unfold_locales (metis assms(2) m_ac(2))+ qed lemma (in comm_monoid_cancel) is_comm_monoid_cancel: "comm_monoid_cancel G" by intro_locales sublocale comm_group \ comm_monoid_cancel .. subsection \Products of Units in Monoids\ lemma (in monoid) prod_unit_l: assumes abunit[simp]: "a \ b \ Units G" and aunit[simp]: "a \ Units G" and carr[simp]: "a \ carrier G" "b \ carrier G" shows "b \ Units G" proof - have c: "inv (a \ b) \ a \ carrier G" by simp have "(inv (a \ b) \ a) \ b = inv (a \ b) \ (a \ b)" by (simp add: m_assoc) also have "\ = \" by simp finally have li: "(inv (a \ b) \ a) \ b = \" . have "\ = inv a \ a" by (simp add: Units_l_inv[symmetric]) also have "\ = inv a \ \ \ a" by simp also have "\ = inv a \ ((a \ b) \ inv (a \ b)) \ a" by (simp add: Units_r_inv[OF abunit, symmetric] del: Units_r_inv) also have "\ = ((inv a \ a) \ b) \ inv (a \ b) \ a" by (simp add: m_assoc del: Units_l_inv) also have "\ = b \ inv (a \ b) \ a" by simp also have "\ = b \ (inv (a \ b) \ a)" by (simp add: m_assoc) finally have ri: "b \ (inv (a \ b) \ a) = \ " by simp from c li ri show "b \ Units G" by (auto simp: Units_def) qed lemma (in monoid) prod_unit_r: assumes abunit[simp]: "a \ b \ Units G" and bunit[simp]: "b \ Units G" and carr[simp]: "a \ carrier G" "b \ carrier G" shows "a \ Units G" proof - have c: "b \ inv (a \ b) \ carrier G" by simp have "a \ (b \ inv (a \ b)) = (a \ b) \ inv (a \ b)" by (simp add: m_assoc del: Units_r_inv) also have "\ = \" by simp finally have li: "a \ (b \ inv (a \ b)) = \" . have "\ = b \ inv b" by (simp add: Units_r_inv[symmetric]) also have "\ = b \ \ \ inv b" by simp also have "\ = b \ (inv (a \ b) \ (a \ b)) \ inv b" by (simp add: Units_l_inv[OF abunit, symmetric] del: Units_l_inv) also have "\ = (b \ inv (a \ b) \ a) \ (b \ inv b)" by (simp add: m_assoc del: Units_l_inv) also have "\ = b \ inv (a \ b) \ a" by simp finally have ri: "(b \ inv (a \ b)) \ a = \ " by simp from c li ri show "a \ Units G" by (auto simp: Units_def) qed lemma (in comm_monoid) unit_factor: assumes abunit: "a \ b \ Units G" and [simp]: "a \ carrier G" "b \ carrier G" shows "a \ Units G" using abunit[simplified Units_def] proof clarsimp fix i assume [simp]: "i \ carrier G" have carr': "b \ i \ carrier G" by simp have "(b \ i) \ a = (i \ b) \ a" by (simp add: m_comm) also have "\ = i \ (b \ a)" by (simp add: m_assoc) also have "\ = i \ (a \ b)" by (simp add: m_comm) also assume "i \ (a \ b) = \" finally have li': "(b \ i) \ a = \" . have "a \ (b \ i) = a \ b \ i" by (simp add: m_assoc) also assume "a \ b \ i = \" finally have ri': "a \ (b \ i) = \" . from carr' li' ri' show "a \ Units G" by (simp add: Units_def, fast) qed subsection \Divisibility and Association\ subsubsection \Function definitions\ definition factor :: "[_, 'a, 'a] \ bool" (infix "divides\" 65) where "a divides\<^bsub>G\<^esub> b \ (\c\carrier G. b = a \\<^bsub>G\<^esub> c)" definition associated :: "[_, 'a, 'a] \ bool" (infix "\\" 55) where "a \\<^bsub>G\<^esub> b \ a divides\<^bsub>G\<^esub> b \ b divides\<^bsub>G\<^esub> a" abbreviation "division_rel G \ \carrier = carrier G, eq = (\\<^bsub>G\<^esub>), le = (divides\<^bsub>G\<^esub>)\" definition properfactor :: "[_, 'a, 'a] \ bool" where "properfactor G a b \ a divides\<^bsub>G\<^esub> b \ \(b divides\<^bsub>G\<^esub> a)" definition irreducible :: "[_, 'a] \ bool" where "irreducible G a \ a \ Units G \ (\b\carrier G. properfactor G b a \ b \ Units G)" definition prime :: "[_, 'a] \ bool" where "prime G p \ p \ Units G \ (\a\carrier G. \b\carrier G. p divides\<^bsub>G\<^esub> (a \\<^bsub>G\<^esub> b) \ p divides\<^bsub>G\<^esub> a \ p divides\<^bsub>G\<^esub> b)" subsubsection \Divisibility\ lemma dividesI: fixes G (structure) assumes carr: "c \ carrier G" and p: "b = a \ c" shows "a divides b" unfolding factor_def using assms by fast lemma dividesI' [intro]: fixes G (structure) assumes p: "b = a \ c" and carr: "c \ carrier G" shows "a divides b" using assms by (fast intro: dividesI) lemma dividesD: fixes G (structure) assumes "a divides b" shows "\c\carrier G. b = a \ c" using assms unfolding factor_def by fast lemma dividesE [elim]: fixes G (structure) assumes d: "a divides b" and elim: "\c. \b = a \ c; c \ carrier G\ \ P" shows "P" proof - from dividesD[OF d] obtain c where "c \ carrier G" and "b = a \ c" by auto then show P by (elim elim) qed lemma (in monoid) divides_refl[simp, intro!]: assumes carr: "a \ carrier G" shows "a divides a" by (intro dividesI[of "\"]) (simp_all add: carr) lemma (in monoid) divides_trans [trans]: assumes dvds: "a divides b" "b divides c" and acarr: "a \ carrier G" shows "a divides c" using dvds[THEN dividesD] by (blast intro: dividesI m_assoc acarr) lemma (in monoid) divides_mult_lI [intro]: assumes "a divides b" "a \ carrier G" "c \ carrier G" shows "(c \ a) divides (c \ b)" by (metis assms factor_def m_assoc) lemma (in monoid_cancel) divides_mult_l [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "(c \ a) divides (c \ b) = a divides b" proof show "c \ a divides c \ b \ a divides b" using carr monoid.m_assoc monoid_axioms monoid_cancel.l_cancel monoid_cancel_axioms by fastforce show "a divides b \ c \ a divides c \ b" using carr(1) carr(3) by blast qed lemma (in comm_monoid) divides_mult_rI [intro]: assumes ab: "a divides b" and carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "(a \ c) divides (b \ c)" using carr ab by (metis divides_mult_lI m_comm) lemma (in comm_monoid_cancel) divides_mult_r [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "(a \ c) divides (b \ c) = a divides b" using carr by (simp add: m_comm[of a c] m_comm[of b c]) lemma (in monoid) divides_prod_r: assumes ab: "a divides b" and carr: "a \ carrier G" "c \ carrier G" shows "a divides (b \ c)" using ab carr by (fast intro: m_assoc) lemma (in comm_monoid) divides_prod_l: assumes "a \ carrier G" "b \ carrier G" "c \ carrier G" "a divides b" shows "a divides (c \ b)" using assms by (simp add: divides_prod_r m_comm) lemma (in monoid) unit_divides: assumes uunit: "u \ Units G" and acarr: "a \ carrier G" shows "u divides a" proof (intro dividesI[of "(inv u) \ a"], fast intro: uunit acarr) from uunit acarr have xcarr: "inv u \ a \ carrier G" by fast from uunit acarr have "u \ (inv u \ a) = (u \ inv u) \ a" by (fast intro: m_assoc[symmetric]) also have "\ = \ \ a" by (simp add: Units_r_inv[OF uunit]) also from acarr have "\ = a" by simp finally show "a = u \ (inv u \ a)" .. qed lemma (in comm_monoid) divides_unit: assumes udvd: "a divides u" and carr: "a \ carrier G" "u \ Units G" shows "a \ Units G" using udvd carr by (blast intro: unit_factor) lemma (in comm_monoid) Unit_eq_dividesone: assumes ucarr: "u \ carrier G" shows "u \ Units G = u divides \" using ucarr by (fast dest: divides_unit intro: unit_divides) subsubsection \Association\ lemma associatedI: fixes G (structure) assumes "a divides b" "b divides a" shows "a \ b" using assms by (simp add: associated_def) lemma (in monoid) associatedI2: assumes uunit[simp]: "u \ Units G" and a: "a = b \ u" and bcarr: "b \ carrier G" shows "a \ b" using uunit bcarr unfolding a apply (intro associatedI) apply (metis Units_closed divides_mult_lI one_closed r_one unit_divides) by blast lemma (in monoid) associatedI2': assumes "a = b \ u" and "u \ Units G" and "b \ carrier G" shows "a \ b" using assms by (intro associatedI2) lemma associatedD: fixes G (structure) assumes "a \ b" shows "a divides b" using assms by (simp add: associated_def) lemma (in monoid_cancel) associatedD2: assumes assoc: "a \ b" and carr: "a \ carrier G" "b \ carrier G" shows "\u\Units G. a = b \ u" using assoc unfolding associated_def proof clarify assume "b divides a" then obtain u where ucarr: "u \ carrier G" and a: "a = b \ u" by (rule dividesE) assume "a divides b" then obtain u' where u'carr: "u' \ carrier G" and b: "b = a \ u'" by (rule dividesE) note carr = carr ucarr u'carr from carr have "a \ \ = a" by simp also have "\ = b \ u" by (simp add: a) also have "\ = a \ u' \ u" by (simp add: b) also from carr have "\ = a \ (u' \ u)" by (simp add: m_assoc) finally have "a \ \ = a \ (u' \ u)" . with carr have u1: "\ = u' \ u" by (fast dest: l_cancel) from carr have "b \ \ = b" by simp also have "\ = a \ u'" by (simp add: b) also have "\ = b \ u \ u'" by (simp add: a) also from carr have "\ = b \ (u \ u')" by (simp add: m_assoc) finally have "b \ \ = b \ (u \ u')" . with carr have u2: "\ = u \ u'" by (fast dest: l_cancel) from u'carr u1[symmetric] u2[symmetric] have "\u'\carrier G. u' \ u = \ \ u \ u' = \" by fast then have "u \ Units G" by (simp add: Units_def ucarr) with ucarr a show "\u\Units G. a = b \ u" by fast qed lemma associatedE: fixes G (structure) assumes assoc: "a \ b" and e: "\a divides b; b divides a\ \ P" shows "P" proof - from assoc have "a divides b" "b divides a" by (simp_all add: associated_def) then show P by (elim e) qed lemma (in monoid_cancel) associatedE2: assumes assoc: "a \ b" and e: "\u. \a = b \ u; u \ Units G\ \ P" and carr: "a \ carrier G" "b \ carrier G" shows "P" proof - from assoc and carr have "\u\Units G. a = b \ u" by (rule associatedD2) then obtain u where "u \ Units G" "a = b \ u" by auto then show P by (elim e) qed lemma (in monoid) associated_refl [simp, intro!]: assumes "a \ carrier G" shows "a \ a" using assms by (fast intro: associatedI) lemma (in monoid) associated_sym [sym]: assumes "a \ b" shows "b \ a" using assms by (iprover intro: associatedI elim: associatedE) lemma (in monoid) associated_trans [trans]: assumes "a \ b" "b \ c" and "a \ carrier G" "c \ carrier G" shows "a \ c" using assms by (iprover intro: associatedI divides_trans elim: associatedE) lemma (in monoid) division_equiv [intro, simp]: "equivalence (division_rel G)" apply unfold_locales apply simp_all apply (metis associated_def) apply (iprover intro: associated_trans) done subsubsection \Division and associativity\ lemmas divides_antisym = associatedI lemma (in monoid) divides_cong_l [trans]: assumes "x \ x'" "x' divides y" "x \ carrier G" shows "x divides y" by (meson assms associatedD divides_trans) lemma (in monoid) divides_cong_r [trans]: assumes "x divides y" "y \ y'" "x \ carrier G" shows "x divides y'" by (meson assms associatedD divides_trans) lemma (in monoid) division_weak_partial_order [simp, intro!]: "weak_partial_order (division_rel G)" apply unfold_locales apply (simp_all add: associated_sym divides_antisym) apply (metis associated_trans) apply (metis divides_trans) by (meson associated_def divides_trans) subsubsection \Multiplication and associativity\ lemma (in monoid) mult_cong_r: assumes "b \ b'" "a \ carrier G" "b \ carrier G" "b' \ carrier G" shows "a \ b \ a \ b'" by (meson assms associated_def divides_mult_lI) lemma (in comm_monoid) mult_cong_l: assumes "a \ a'" "a \ carrier G" "a' \ carrier G" "b \ carrier G" shows "a \ b \ a' \ b" using assms m_comm mult_cong_r by auto lemma (in monoid_cancel) assoc_l_cancel: assumes "a \ carrier G" "b \ carrier G" "b' \ carrier G" "a \ b \ a \ b'" shows "b \ b'" by (meson assms associated_def divides_mult_l) lemma (in comm_monoid_cancel) assoc_r_cancel: assumes "a \ b \ a' \ b" "a \ carrier G" "a' \ carrier G" "b \ carrier G" shows "a \ a'" using assms assoc_l_cancel m_comm by presburger subsubsection \Units\ lemma (in monoid_cancel) assoc_unit_l [trans]: assumes "a \ b" and "b \ Units G" and "a \ carrier G" shows "a \ Units G" using assms by (fast elim: associatedE2) lemma (in monoid_cancel) assoc_unit_r [trans]: assumes aunit: "a \ Units G" and asc: "a \ b" and bcarr: "b \ carrier G" shows "b \ Units G" using aunit bcarr associated_sym[OF asc] by (blast intro: assoc_unit_l) lemma (in comm_monoid) Units_cong: assumes aunit: "a \ Units G" and asc: "a \ b" and bcarr: "b \ carrier G" shows "b \ Units G" using assms by (blast intro: divides_unit elim: associatedE) lemma (in monoid) Units_assoc: assumes units: "a \ Units G" "b \ Units G" shows "a \ b" using units by (fast intro: associatedI unit_divides) lemma (in monoid) Units_are_ones: "Units G {.=}\<^bsub>(division_rel G)\<^esub> {\}" proof - have "a .\\<^bsub>division_rel G\<^esub> {\}" if "a \ Units G" for a proof - have "a \ \" by (rule associatedI) (simp_all add: Units_closed that unit_divides) then show ?thesis by (simp add: elem_def) qed moreover have "\ .\\<^bsub>division_rel G\<^esub> Units G" by (simp add: equivalence.mem_imp_elem) ultimately show ?thesis by (auto simp: set_eq_def) qed lemma (in comm_monoid) Units_Lower: "Units G = Lower (division_rel G) (carrier G)" apply (auto simp add: Units_def Lower_def) apply (metis Units_one_closed unit_divides unit_factor) apply (metis Unit_eq_dividesone Units_r_inv_ex m_ac(2) one_closed) done lemma (in monoid_cancel) associated_iff: assumes "a \ carrier G" "b \ carrier G" shows "a \ b \ (\c \ Units G. a = b \ c)" using assms associatedI2' associatedD2 by auto subsubsection \Proper factors\ lemma properfactorI: fixes G (structure) assumes "a divides b" and "\(b divides a)" shows "properfactor G a b" using assms unfolding properfactor_def by simp lemma properfactorI2: fixes G (structure) assumes advdb: "a divides b" and neq: "\(a \ b)" shows "properfactor G a b" proof (rule properfactorI, rule advdb, rule notI) assume "b divides a" with advdb have "a \ b" by (rule associatedI) with neq show "False" by fast qed lemma (in comm_monoid_cancel) properfactorI3: assumes p: "p = a \ b" and nunit: "b \ Units G" and carr: "a \ carrier G" "b \ carrier G" shows "properfactor G a p" unfolding p using carr apply (intro properfactorI, fast) proof (clarsimp, elim dividesE) fix c assume ccarr: "c \ carrier G" note [simp] = carr ccarr have "a \ \ = a" by simp also assume "a = a \ b \ c" also have "\ = a \ (b \ c)" by (simp add: m_assoc) finally have "a \ \ = a \ (b \ c)" . then have rinv: "\ = b \ c" by (intro l_cancel[of "a" "\" "b \ c"], simp+) also have "\ = c \ b" by (simp add: m_comm) finally have linv: "\ = c \ b" . from ccarr linv[symmetric] rinv[symmetric] have "b \ Units G" unfolding Units_def by fastforce with nunit show False .. qed lemma properfactorE: fixes G (structure) assumes pf: "properfactor G a b" and r: "\a divides b; \(b divides a)\ \ P" shows "P" using pf unfolding properfactor_def by (fast intro: r) lemma properfactorE2: fixes G (structure) assumes pf: "properfactor G a b" and elim: "\a divides b; \(a \ b)\ \ P" shows "P" using pf unfolding properfactor_def by (fast elim: elim associatedE) lemma (in monoid) properfactor_unitE: assumes uunit: "u \ Units G" and pf: "properfactor G a u" and acarr: "a \ carrier G" shows "P" using pf unit_divides[OF uunit acarr] by (fast elim: properfactorE) lemma (in monoid) properfactor_divides: assumes pf: "properfactor G a b" shows "a divides b" using pf by (elim properfactorE) lemma (in monoid) properfactor_trans1 [trans]: assumes "a divides b" "properfactor G b c" "a \ carrier G" "c \ carrier G" shows "properfactor G a c" by (meson divides_trans properfactorE properfactorI assms) lemma (in monoid) properfactor_trans2 [trans]: assumes "properfactor G a b" "b divides c" "a \ carrier G" "b \ carrier G" shows "properfactor G a c" by (meson divides_trans properfactorE properfactorI assms) lemma properfactor_lless: fixes G (structure) shows "properfactor G = lless (division_rel G)" by (force simp: lless_def properfactor_def associated_def) lemma (in monoid) properfactor_cong_l [trans]: assumes x'x: "x' \ x" and pf: "properfactor G x y" and carr: "x \ carrier G" "x' \ carrier G" "y \ carrier G" shows "properfactor G x' y" using pf unfolding properfactor_lless proof - interpret weak_partial_order "division_rel G" .. from x'x have "x' .=\<^bsub>division_rel G\<^esub> x" by simp also assume "x \\<^bsub>division_rel G\<^esub> y" finally show "x' \\<^bsub>division_rel G\<^esub> y" by (simp add: carr) qed lemma (in monoid) properfactor_cong_r [trans]: assumes pf: "properfactor G x y" and yy': "y \ y'" and carr: "x \ carrier G" "y \ carrier G" "y' \ carrier G" shows "properfactor G x y'" using pf unfolding properfactor_lless proof - interpret weak_partial_order "division_rel G" .. assume "x \\<^bsub>division_rel G\<^esub> y" also from yy' have "y .=\<^bsub>division_rel G\<^esub> y'" by simp finally show "x \\<^bsub>division_rel G\<^esub> y'" by (simp add: carr) qed lemma (in monoid_cancel) properfactor_mult_lI [intro]: assumes ab: "properfactor G a b" and carr: "a \ carrier G" "c \ carrier G" shows "properfactor G (c \ a) (c \ b)" using ab carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in monoid_cancel) properfactor_mult_l [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G (c \ a) (c \ b) = properfactor G a b" using carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in comm_monoid_cancel) properfactor_mult_rI [intro]: assumes ab: "properfactor G a b" and carr: "a \ carrier G" "c \ carrier G" shows "properfactor G (a \ c) (b \ c)" using ab carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in comm_monoid_cancel) properfactor_mult_r [simp]: assumes carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G (a \ c) (b \ c) = properfactor G a b" using carr by (fastforce elim: properfactorE intro: properfactorI) lemma (in monoid) properfactor_prod_r: assumes ab: "properfactor G a b" and carr[simp]: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G a (b \ c)" by (intro properfactor_trans2[OF ab] divides_prod_r) simp_all lemma (in comm_monoid) properfactor_prod_l: assumes ab: "properfactor G a b" and carr[simp]: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "properfactor G a (c \ b)" by (intro properfactor_trans2[OF ab] divides_prod_l) simp_all subsection \Irreducible Elements and Primes\ subsubsection \Irreducible elements\ lemma irreducibleI: fixes G (structure) assumes "a \ Units G" and "\b. \b \ carrier G; properfactor G b a\ \ b \ Units G" shows "irreducible G a" using assms unfolding irreducible_def by blast lemma irreducibleE: fixes G (structure) assumes irr: "irreducible G a" and elim: "\a \ Units G; \b. b \ carrier G \ properfactor G b a \ b \ Units G\ \ P" shows "P" using assms unfolding irreducible_def by blast lemma irreducibleD: fixes G (structure) assumes irr: "irreducible G a" and pf: "properfactor G b a" and bcarr: "b \ carrier G" shows "b \ Units G" using assms by (fast elim: irreducibleE) lemma (in monoid_cancel) irreducible_cong [trans]: assumes "irreducible G a" "a \ a'" "a \ carrier G" "a' \ carrier G" shows "irreducible G a'" proof - have "a' divides a" by (meson \a \ a'\ associated_def) then show ?thesis by (metis (no_types) assms assoc_unit_l irreducibleE irreducibleI monoid.properfactor_trans2 monoid_axioms) qed lemma (in monoid) irreducible_prod_rI: assumes "irreducible G a" "b \ Units G" "a \ carrier G" "b \ carrier G" shows "irreducible G (a \ b)" using assms by (metis (no_types, lifting) associatedI2' irreducible_def monoid.m_closed monoid_axioms prod_unit_r properfactor_cong_r) lemma (in comm_monoid) irreducible_prod_lI: assumes birr: "irreducible G b" and aunit: "a \ Units G" and carr [simp]: "a \ carrier G" "b \ carrier G" shows "irreducible G (a \ b)" by (metis aunit birr carr irreducible_prod_rI m_comm) lemma (in comm_monoid_cancel) irreducible_prodE [elim]: assumes irr: "irreducible G (a \ b)" and carr[simp]: "a \ carrier G" "b \ carrier G" and e1: "\irreducible G a; b \ Units G\ \ P" and e2: "\a \ Units G; irreducible G b\ \ P" shows P using irr proof (elim irreducibleE) assume abnunit: "a \ b \ Units G" and isunit[rule_format]: "\ba. ba \ carrier G \ properfactor G ba (a \ b) \ ba \ Units G" show P proof (cases "a \ Units G") case aunit: True have "irreducible G b" proof (rule irreducibleI, rule notI) assume "b \ Units G" with aunit have "(a \ b) \ Units G" by fast with abnunit show "False" .. next fix c assume ccarr: "c \ carrier G" and "properfactor G c b" then have "properfactor G c (a \ b)" by (simp add: properfactor_prod_l[of c b a]) with ccarr show "c \ Units G" by (fast intro: isunit) qed with aunit show "P" by (rule e2) next case anunit: False with carr have "properfactor G b (b \ a)" by (fast intro: properfactorI3) then have bf: "properfactor G b (a \ b)" by (subst m_comm[of a b], simp+) then have bunit: "b \ Units G" by (intro isunit, simp) have "irreducible G a" proof (rule irreducibleI, rule notI) assume "a \ Units G" with bunit have "(a \ b) \ Units G" by fast with abnunit show "False" .. next fix c assume ccarr: "c \ carrier G" and "properfactor G c a" then have "properfactor G c (a \ b)" by (simp add: properfactor_prod_r[of c a b]) with ccarr show "c \ Units G" by (fast intro: isunit) qed from this bunit show "P" by (rule e1) qed qed lemma divides_irreducible_condition: assumes "irreducible G r" and "a \ carrier G" shows "a divides\<^bsub>G\<^esub> r \ a \ Units G \ a \\<^bsub>G\<^esub> r" using assms unfolding irreducible_def properfactor_def associated_def by (cases "r divides\<^bsub>G\<^esub> a", auto) subsubsection \Prime elements\ lemma primeI: fixes G (structure) assumes "p \ Units G" and "\a b. \a \ carrier G; b \ carrier G; p divides (a \ b)\ \ p divides a \ p divides b" shows "prime G p" using assms unfolding prime_def by blast lemma primeE: fixes G (structure) assumes pprime: "prime G p" and e: "\p \ Units G; \a\carrier G. \b\carrier G. p divides a \ b \ p divides a \ p divides b\ \ P" shows "P" using pprime unfolding prime_def by (blast dest: e) lemma (in comm_monoid_cancel) prime_divides: assumes carr: "a \ carrier G" "b \ carrier G" and pprime: "prime G p" and pdvd: "p divides a \ b" shows "p divides a \ p divides b" using assms by (blast elim: primeE) lemma (in monoid_cancel) prime_cong [trans]: assumes "prime G p" and pp': "p \ p'" "p \ carrier G" "p' \ carrier G" shows "prime G p'" using assms by (auto simp: prime_def assoc_unit_l) (metis pp' associated_sym divides_cong_l) lemma (in comm_monoid_cancel) prime_irreducible: \<^marker>\contributor \Paulo Emílio de Vilhena\\ assumes "prime G p" shows "irreducible G p" proof (rule irreducibleI) show "p \ Units G" using assms unfolding prime_def by simp next fix b assume A: "b \ carrier G" "properfactor G b p" then obtain c where c: "c \ carrier G" "p = b \ c" unfolding properfactor_def factor_def by auto hence "p divides c" using A assms unfolding prime_def properfactor_def by auto then obtain b' where b': "b' \ carrier G" "c = p \ b'" unfolding factor_def by auto hence "\ = b \ b'" by (metis A(1) l_cancel m_closed m_lcomm one_closed r_one c) thus "b \ Units G" using A(1) Units_one_closed b'(1) unit_factor by presburger qed lemma (in comm_monoid_cancel) prime_pow_divides_iff: assumes "p \ carrier G" "a \ carrier G" "b \ carrier G" and "prime G p" and "\ (p divides a)" shows "(p [^] (n :: nat)) divides (a \ b) \ (p [^] n) divides b" proof assume "(p [^] n) divides b" thus "(p [^] n) divides (a \ b)" using divides_prod_l[of "p [^] n" b a] assms by simp next assume "(p [^] n) divides (a \ b)" thus "(p [^] n) divides b" proof (induction n) case 0 with \b \ carrier G\ show ?case by (simp add: unit_divides) next case (Suc n) hence "(p [^] n) divides (a \ b)" and "(p [^] n) divides b" using assms(1) divides_prod_r by auto with \(p [^] (Suc n)) divides (a \ b)\ obtain c d where c: "c \ carrier G" and "b = (p [^] n) \ c" and d: "d \ carrier G" and "a \ b = (p [^] (Suc n)) \ d" using assms by blast hence "(p [^] n) \ (a \ c) = (p [^] n) \ (p \ d)" using assms by (simp add: m_assoc m_lcomm) hence "a \ c = p \ d" using c d assms(1) assms(2) l_cancel by blast with \\ (p divides a)\ and \prime G p\ have "p divides c" by (metis assms(2) c d dividesI' prime_divides) with \b = (p [^] n) \ c\ show ?case using assms(1) c by simp qed qed subsection \Factorization and Factorial Monoids\ subsubsection \Function definitions\ definition factors :: "('a, _) monoid_scheme \ 'a list \ 'a \ bool" where "factors G fs a \ (\x \ (set fs). irreducible G x) \ foldr (\\<^bsub>G\<^esub>) fs \\<^bsub>G\<^esub> = a" definition wfactors ::"('a, _) monoid_scheme \ 'a list \ 'a \ bool" where "wfactors G fs a \ (\x \ (set fs). irreducible G x) \ foldr (\\<^bsub>G\<^esub>) fs \\<^bsub>G\<^esub> \\<^bsub>G\<^esub> a" abbreviation list_assoc :: "('a, _) monoid_scheme \ 'a list \ 'a list \ bool" (infix "[\]\" 44) where "list_assoc G \ list_all2 (\\<^bsub>G\<^esub>)" definition essentially_equal :: "('a, _) monoid_scheme \ 'a list \ 'a list \ bool" where "essentially_equal G fs1 fs2 \ (\fs1'. fs1 <~~> fs1' \ fs1' [\]\<^bsub>G\<^esub> fs2)" locale factorial_monoid = comm_monoid_cancel + assumes factors_exist: "\a \ carrier G; a \ Units G\ \ \fs. set fs \ carrier G \ factors G fs a" and factors_unique: "\factors G fs a; factors G fs' a; a \ carrier G; a \ Units G; set fs \ carrier G; set fs' \ carrier G\ \ essentially_equal G fs fs'" subsubsection \Comparing lists of elements\ text \Association on lists\ lemma (in monoid) listassoc_refl [simp, intro]: assumes "set as \ carrier G" shows "as [\] as" using assms by (induct as) simp_all lemma (in monoid) listassoc_sym [sym]: assumes "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "bs [\] as" using assms proof (induction as arbitrary: bs) case Cons then show ?case by (induction bs) (use associated_sym in auto) qed auto lemma (in monoid) listassoc_trans [trans]: assumes "as [\] bs" and "bs [\] cs" and "set as \ carrier G" and "set bs \ carrier G" and "set cs \ carrier G" shows "as [\] cs" using assms apply (simp add: list_all2_conv_all_nth set_conv_nth, safe) by (metis (mono_tags, lifting) associated_trans nth_mem subsetCE) lemma (in monoid_cancel) irrlist_listassoc_cong: assumes "\a\set as. irreducible G a" and "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "\a\set bs. irreducible G a" using assms by (fastforce simp add: list_all2_conv_all_nth set_conv_nth intro: irreducible_cong) text \Permutations\ lemma perm_map [intro]: assumes p: "a <~~> b" shows "map f a <~~> map f b" - using p by (simp add: perm_iff_eq_mset) + using p by simp lemma perm_map_switch: assumes m: "map f a = map f b" and p: "b <~~> c" shows "\d. a <~~> d \ map f d = map f c" proof - from m have \length a = length b\ by (rule map_eq_imp_length_eq) from p have \mset c = mset b\ - by (simp add: perm_iff_eq_mset) + by simp then obtain p where \p permutes {.. \permute_list p b = c\ by (rule mset_eq_permutation) with \length a = length b\ have \p permutes {.. by simp moreover define d where \d = permute_list p a\ ultimately have \mset a = mset d\ \map f d = map f c\ using m \p permutes {.. \permute_list p b = c\ by (auto simp flip: permute_list_map) then show ?thesis - by (auto simp add: perm_iff_eq_mset) + by auto qed lemma (in monoid) perm_assoc_switch: assumes a:"as [\] bs" and p: "bs <~~> cs" shows "\bs'. as <~~> bs' \ bs' [\] cs" - using p a -proof (induction bs cs arbitrary: as) - case (swap y x l) - then show ?case - by (metis (no_types, hide_lams) list_all2_Cons2 perm.swap) -next -case (Cons xs ys z) - then show ?case - by (metis list_all2_Cons2 perm.Cons) -next - case (trans xs ys zs) - then show ?case - by (meson perm.trans) -qed auto +proof - + from p have \mset cs = mset bs\ + by simp + then obtain p where \p permutes {.. \permute_list p bs = cs\ + by (rule mset_eq_permutation) + moreover define bs' where \bs' = permute_list p as\ + ultimately have \as <~~> bs'\ and \bs' [\] cs\ + using a by (auto simp add: list_all2_permute_list_iff list_all2_lengthD) + then show ?thesis by blast +qed lemma (in monoid) perm_assoc_switch_r: assumes p: "as <~~> bs" and a:"bs [\] cs" shows "\bs'. as [\] bs' \ bs' <~~> cs" - using p a -proof (induction as bs arbitrary: cs) - case Nil - then show ?case - by auto -next - case (swap y x l) - then show ?case - by (metis (no_types, hide_lams) list_all2_Cons1 perm.swap) -next - case (Cons xs ys z) - then show ?case - by (metis list_all2_Cons1 perm.Cons) -next - case (trans xs ys zs) - then show ?case - by (blast intro: elim: ) -qed + using a p by (rule list_all2_reorder_left_invariance) declare perm_sym [sym] lemma perm_setP: assumes perm: "as <~~> bs" and as: "P (set as)" shows "P (set bs)" -proof - - from perm have "mset as = mset bs" - by (simp add: mset_eq_perm) - then have "set as = set bs" - by (rule mset_eq_setD) - with as show "P (set bs)" - by simp -qed + using assms by (metis set_mset_mset) lemmas (in monoid) perm_closed = perm_setP[of _ _ "\as. as \ carrier G"] lemmas (in monoid) irrlist_perm_cong = perm_setP[of _ _ "\as. \a\as. irreducible G a"] text \Essentially equal factorizations\ lemma (in monoid) essentially_equalI: assumes ex: "fs1 <~~> fs1'" "fs1' [\] fs2" shows "essentially_equal G fs1 fs2" using ex unfolding essentially_equal_def by fast lemma (in monoid) essentially_equalE: assumes ee: "essentially_equal G fs1 fs2" and e: "\fs1'. \fs1 <~~> fs1'; fs1' [\] fs2\ \ P" shows "P" using ee unfolding essentially_equal_def by (fast intro: e) lemma (in monoid) ee_refl [simp,intro]: assumes carr: "set as \ carrier G" shows "essentially_equal G as as" using carr by (fast intro: essentially_equalI) lemma (in monoid) ee_sym [sym]: assumes ee: "essentially_equal G as bs" and carr: "set as \ carrier G" "set bs \ carrier G" shows "essentially_equal G bs as" using ee proof (elim essentially_equalE) fix fs assume "as <~~> fs" "fs [\] bs" from perm_assoc_switch_r [OF this] obtain fs' where a: "as [\] fs'" and p: "fs' <~~> bs" by blast from p have "bs <~~> fs'" by (rule perm_sym) with a[symmetric] carr show ?thesis by (iprover intro: essentially_equalI perm_closed) qed lemma (in monoid) ee_trans [trans]: assumes ab: "essentially_equal G as bs" and bc: "essentially_equal G bs cs" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" and cscarr: "set cs \ carrier G" shows "essentially_equal G as cs" using ab bc proof (elim essentially_equalE) fix abs bcs assume "abs [\] bs" and pb: "bs <~~> bcs" from perm_assoc_switch [OF this] obtain bs' where p: "abs <~~> bs'" and a: "bs' [\] bcs" by blast assume "as <~~> abs" - with p have pp: "as <~~> bs'" by fast + with p have pp: "as <~~> bs'" by simp from pp ascarr have c1: "set bs' \ carrier G" by (rule perm_closed) from pb bscarr have c2: "set bcs \ carrier G" by (rule perm_closed) assume "bcs [\] cs" then have "bs' [\] cs" using a c1 c2 cscarr listassoc_trans by blast with pp show ?thesis by (rule essentially_equalI) qed subsubsection \Properties of lists of elements\ text \Multiplication of factors in a list\ lemma (in monoid) multlist_closed [simp, intro]: assumes ascarr: "set fs \ carrier G" shows "foldr (\) fs \ \ carrier G" using ascarr by (induct fs) simp_all lemma (in comm_monoid) multlist_dividesI: assumes "f \ set fs" and "set fs \ carrier G" shows "f divides (foldr (\) fs \)" using assms proof (induction fs) case (Cons a fs) then have f: "f \ carrier G" by blast show ?case using Cons.IH Cons.prems(1) Cons.prems(2) divides_prod_l f by auto qed auto lemma (in comm_monoid_cancel) multlist_listassoc_cong: assumes "fs [\] fs'" and "set fs \ carrier G" and "set fs' \ carrier G" shows "foldr (\) fs \ \ foldr (\) fs' \" using assms proof (induct fs arbitrary: fs') case (Cons a as fs') then show ?case proof (induction fs') case (Cons b bs) then have p: "a \ foldr (\) as \ \ b \ foldr (\) as \" by (simp add: mult_cong_l) then have "foldr (\) as \ \ foldr (\) bs \" using Cons by auto with Cons have "b \ foldr (\) as \ \ b \ foldr (\) bs \" by (simp add: mult_cong_r) then show ?case using Cons.prems(3) Cons.prems(4) monoid.associated_trans monoid_axioms p by force qed auto qed auto lemma (in comm_monoid) multlist_perm_cong: assumes prm: "as <~~> bs" and ascarr: "set as \ carrier G" shows "foldr (\) as \ = foldr (\) bs \" - using prm ascarr -proof induction - case (swap y x l) then show ?case - by (simp add: m_lcomm) -next - case (trans xs ys zs) then show ?case - using perm_closed by auto -qed auto +proof - + from prm have \mset (rev as) = mset (rev bs)\ + by simp + moreover note one_closed + ultimately have \fold (\) (rev as) \ = fold (\) (rev bs) \\ + by (rule fold_permuted_eq) (use ascarr in \auto intro: m_lcomm\) + then show ?thesis + by (simp add: foldr_conv_fold) +qed lemma (in comm_monoid_cancel) multlist_ee_cong: assumes "essentially_equal G fs fs'" and "set fs \ carrier G" and "set fs' \ carrier G" shows "foldr (\) fs \ \ foldr (\) fs' \" using assms by (metis essentially_equal_def multlist_listassoc_cong multlist_perm_cong perm_closed) subsubsection \Factorization in irreducible elements\ lemma wfactorsI: fixes G (structure) assumes "\f\set fs. irreducible G f" and "foldr (\) fs \ \ a" shows "wfactors G fs a" using assms unfolding wfactors_def by simp lemma wfactorsE: fixes G (structure) assumes wf: "wfactors G fs a" and e: "\\f\set fs. irreducible G f; foldr (\) fs \ \ a\ \ P" shows "P" using wf unfolding wfactors_def by (fast dest: e) lemma (in monoid) factorsI: assumes "\f\set fs. irreducible G f" and "foldr (\) fs \ = a" shows "factors G fs a" using assms unfolding factors_def by simp lemma factorsE: fixes G (structure) assumes f: "factors G fs a" and e: "\\f\set fs. irreducible G f; foldr (\) fs \ = a\ \ P" shows "P" using f unfolding factors_def by (simp add: e) lemma (in monoid) factors_wfactors: assumes "factors G as a" and "set as \ carrier G" shows "wfactors G as a" using assms by (blast elim: factorsE intro: wfactorsI) lemma (in monoid) wfactors_factors: assumes "wfactors G as a" and "set as \ carrier G" shows "\a'. factors G as a' \ a' \ a" using assms by (blast elim: wfactorsE intro: factorsI) lemma (in monoid) factors_closed [dest]: assumes "factors G fs a" and "set fs \ carrier G" shows "a \ carrier G" using assms by (elim factorsE, clarsimp) lemma (in monoid) nunit_factors: assumes anunit: "a \ Units G" and fs: "factors G as a" shows "length as > 0" proof - from anunit Units_one_closed have "a \ \" by auto with fs show ?thesis by (auto elim: factorsE) qed lemma (in monoid) unit_wfactors [simp]: assumes aunit: "a \ Units G" shows "wfactors G [] a" using aunit by (intro wfactorsI) (simp, simp add: Units_assoc) lemma (in comm_monoid_cancel) unit_wfactors_empty: assumes aunit: "a \ Units G" and wf: "wfactors G fs a" and carr[simp]: "set fs \ carrier G" shows "fs = []" proof (cases fs) case fs: (Cons f fs') from carr have fcarr[simp]: "f \ carrier G" and carr'[simp]: "set fs' \ carrier G" by (simp_all add: fs) from fs wf have "irreducible G f" by (simp add: wfactors_def) then have fnunit: "f \ Units G" by (fast elim: irreducibleE) from fs wf have a: "f \ foldr (\) fs' \ \ a" by (simp add: wfactors_def) note aunit also from fs wf have a: "f \ foldr (\) fs' \ \ a" by (simp add: wfactors_def) have "a \ f \ foldr (\) fs' \" by (simp add: Units_closed[OF aunit] a[symmetric]) finally have "f \ foldr (\) fs' \ \ Units G" by simp then have "f \ Units G" by (intro unit_factor[of f], simp+) with fnunit show ?thesis by contradiction qed text \Comparing wfactors\ lemma (in comm_monoid_cancel) wfactors_listassoc_cong_l: assumes fact: "wfactors G fs a" and asc: "fs [\] fs'" and carr: "a \ carrier G" "set fs \ carrier G" "set fs' \ carrier G" shows "wfactors G fs' a" proof - { from asc[symmetric] have "foldr (\) fs' \ \ foldr (\) fs \" by (simp add: multlist_listassoc_cong carr) also assume "foldr (\) fs \ \ a" finally have "foldr (\) fs' \ \ a" by (simp add: carr) } then show ?thesis using fact by (meson asc carr(2) carr(3) irrlist_listassoc_cong wfactors_def) qed lemma (in comm_monoid) wfactors_perm_cong_l: assumes "wfactors G fs a" and "fs <~~> fs'" and "set fs \ carrier G" shows "wfactors G fs' a" using assms irrlist_perm_cong multlist_perm_cong wfactors_def by fastforce lemma (in comm_monoid_cancel) wfactors_ee_cong_l [trans]: assumes ee: "essentially_equal G as bs" and bfs: "wfactors G bs b" and carr: "b \ carrier G" "set as \ carrier G" "set bs \ carrier G" shows "wfactors G as b" using ee proof (elim essentially_equalE) fix fs assume prm: "as <~~> fs" - with carr have fscarr: "set fs \ carrier G" by (simp add: perm_closed) + with carr have fscarr: "set fs \ carrier G" + using perm_closed by blast note bfs also assume [symmetric]: "fs [\] bs" also (wfactors_listassoc_cong_l) - note prm[symmetric] + have \mset fs = mset as\ using prm by simp finally (wfactors_perm_cong_l) show "wfactors G as b" by (simp add: carr fscarr) qed lemma (in monoid) wfactors_cong_r [trans]: assumes fac: "wfactors G fs a" and aa': "a \ a'" and carr[simp]: "a \ carrier G" "a' \ carrier G" "set fs \ carrier G" shows "wfactors G fs a'" using fac proof (elim wfactorsE, intro wfactorsI) assume "foldr (\) fs \ \ a" also note aa' finally show "foldr (\) fs \ \ a'" by simp qed subsubsection \Essentially equal factorizations\ lemma (in comm_monoid_cancel) unitfactor_ee: assumes uunit: "u \ Units G" and carr: "set as \ carrier G" shows "essentially_equal G (as[0 := (as!0 \ u)]) as" (is "essentially_equal G ?as' as") proof - have "as[0 := as ! 0 \ u] [\] as" proof (cases as) case (Cons a as') then show ?thesis using associatedI2 carr uunit by auto qed auto then show ?thesis using essentially_equal_def by blast qed lemma (in comm_monoid_cancel) factors_cong_unit: assumes u: "u \ Units G" and a: "a \ Units G" and afs: "factors G as a" and ascarr: "set as \ carrier G" shows "factors G (as[0 := (as!0 \ u)]) (a \ u)" (is "factors G ?as' ?a'") proof (cases as) case Nil then show ?thesis using afs a nunit_factors by auto next case (Cons b bs) have *: "\f\set as. irreducible G f" "foldr (\) as \ = a" using afs by (auto simp: factors_def) show ?thesis proof (intro factorsI) show "foldr (\) (as[0 := as ! 0 \ u]) \ = a \ u" using Cons u ascarr * by (auto simp add: m_ac Units_closed) show "\f\set (as[0 := as ! 0 \ u]). irreducible G f" using Cons u ascarr * by (force intro: irreducible_prod_rI) qed qed lemma (in comm_monoid) perm_wfactorsD: assumes prm: "as <~~> bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and [simp]: "a \ carrier G" "b \ carrier G" and ascarr [simp]: "set as \ carrier G" shows "a \ b" using afs bfs proof (elim wfactorsE) from prm have [simp]: "set bs \ carrier G" by (simp add: perm_closed) assume "foldr (\) as \ \ a" then have "a \ foldr (\) as \" by (simp add: associated_sym) also from prm have "foldr (\) as \ = foldr (\) bs \" by (rule multlist_perm_cong, simp) also assume "foldr (\) bs \ \ b" finally show "a \ b" by simp qed lemma (in comm_monoid_cancel) listassoc_wfactorsD: assumes assoc: "as [\] bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and [simp]: "a \ carrier G" "b \ carrier G" and [simp]: "set as \ carrier G" "set bs \ carrier G" shows "a \ b" using afs bfs proof (elim wfactorsE) assume "foldr (\) as \ \ a" then have "a \ foldr (\) as \" by (simp add: associated_sym) also from assoc have "foldr (\) as \ \ foldr (\) bs \" by (rule multlist_listassoc_cong, simp+) also assume "foldr (\) bs \ \ b" finally show "a \ b" by simp qed lemma (in comm_monoid_cancel) ee_wfactorsD: assumes ee: "essentially_equal G as bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and [simp]: "a \ carrier G" "b \ carrier G" and ascarr[simp]: "set as \ carrier G" and bscarr[simp]: "set bs \ carrier G" shows "a \ b" using ee proof (elim essentially_equalE) fix fs assume prm: "as <~~> fs" then have as'carr[simp]: "set fs \ carrier G" by (simp add: perm_closed) from afs prm have afs': "wfactors G fs a" by (rule wfactors_perm_cong_l) simp assume "fs [\] bs" from this afs' bfs show "a \ b" by (rule listassoc_wfactorsD) simp_all qed lemma (in comm_monoid_cancel) ee_factorsD: assumes ee: "essentially_equal G as bs" and afs: "factors G as a" and bfs:"factors G bs b" and "set as \ carrier G" "set bs \ carrier G" shows "a \ b" using assms by (blast intro: factors_wfactors dest: ee_wfactorsD) lemma (in factorial_monoid) ee_factorsI: assumes ab: "a \ b" and afs: "factors G as a" and anunit: "a \ Units G" and bfs: "factors G bs b" and bnunit: "b \ Units G" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "essentially_equal G as bs" proof - note carr[simp] = factors_closed[OF afs ascarr] ascarr[THEN subsetD] factors_closed[OF bfs bscarr] bscarr[THEN subsetD] from ab carr obtain u where uunit: "u \ Units G" and a: "a = b \ u" by (elim associatedE2) from uunit bscarr have ee: "essentially_equal G (bs[0 := (bs!0 \ u)]) bs" (is "essentially_equal G ?bs' bs") by (rule unitfactor_ee) from bscarr uunit have bs'carr: "set ?bs' \ carrier G" by (cases bs) (simp_all add: Units_closed) from uunit bnunit bfs bscarr have fac: "factors G ?bs' (b \ u)" by (rule factors_cong_unit) from afs fac[simplified a[symmetric]] ascarr bs'carr anunit have "essentially_equal G as ?bs'" by (blast intro: factors_unique) also note ee finally show "essentially_equal G as bs" by (simp add: ascarr bscarr bs'carr) qed lemma (in factorial_monoid) ee_wfactorsI: assumes asc: "a \ b" and asf: "wfactors G as a" and bsf: "wfactors G bs b" and acarr[simp]: "a \ carrier G" and bcarr[simp]: "b \ carrier G" and ascarr[simp]: "set as \ carrier G" and bscarr[simp]: "set bs \ carrier G" shows "essentially_equal G as bs" using assms proof (cases "a \ Units G") case aunit: True also note asc finally have bunit: "b \ Units G" by simp from aunit asf ascarr have e: "as = []" by (rule unit_wfactors_empty) from bunit bsf bscarr have e': "bs = []" by (rule unit_wfactors_empty) have "essentially_equal G [] []" by (fast intro: essentially_equalI) then show ?thesis by (simp add: e e') next case anunit: False have bnunit: "b \ Units G" proof clarify assume "b \ Units G" also note asc[symmetric] finally have "a \ Units G" by simp with anunit show False .. qed from wfactors_factors[OF asf ascarr] obtain a' where fa': "factors G as a'" and a': "a' \ a" by blast from fa' ascarr have a'carr[simp]: "a' \ carrier G" by fast have a'nunit: "a' \ Units G" proof clarify assume "a' \ Units G" also note a' finally have "a \ Units G" by simp with anunit show "False" .. qed from wfactors_factors[OF bsf bscarr] obtain b' where fb': "factors G bs b'" and b': "b' \ b" by blast from fb' bscarr have b'carr[simp]: "b' \ carrier G" by fast have b'nunit: "b' \ Units G" proof clarify assume "b' \ Units G" also note b' finally have "b \ Units G" by simp with bnunit show False .. qed note a' also note asc also note b'[symmetric] finally have "a' \ b'" by simp from this fa' a'nunit fb' b'nunit ascarr bscarr show "essentially_equal G as bs" by (rule ee_factorsI) qed lemma (in factorial_monoid) ee_wfactors: assumes asf: "wfactors G as a" and bsf: "wfactors G bs b" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows asc: "a \ b = essentially_equal G as bs" using assms by (fast intro: ee_wfactorsI ee_wfactorsD) lemma (in factorial_monoid) wfactors_exist [intro, simp]: assumes acarr[simp]: "a \ carrier G" shows "\fs. set fs \ carrier G \ wfactors G fs a" proof (cases "a \ Units G") case True then have "wfactors G [] a" by (rule unit_wfactors) then show ?thesis by (intro exI) force next case False with factors_exist [OF acarr] obtain fs where fscarr: "set fs \ carrier G" and f: "factors G fs a" by blast from f have "wfactors G fs a" by (rule factors_wfactors) fact with fscarr show ?thesis by fast qed lemma (in monoid) wfactors_prod_exists [intro, simp]: assumes "\a \ set as. irreducible G a" and "set as \ carrier G" shows "\a. a \ carrier G \ wfactors G as a" unfolding wfactors_def using assms by blast lemma (in factorial_monoid) wfactors_unique: assumes "wfactors G fs a" and "wfactors G fs' a" and "a \ carrier G" and "set fs \ carrier G" and "set fs' \ carrier G" shows "essentially_equal G fs fs'" using assms by (fast intro: ee_wfactorsI[of a a]) lemma (in monoid) factors_mult_single: assumes "irreducible G a" and "factors G fb b" and "a \ carrier G" shows "factors G (a # fb) (a \ b)" using assms unfolding factors_def by simp lemma (in monoid_cancel) wfactors_mult_single: assumes f: "irreducible G a" "wfactors G fb b" "a \ carrier G" "b \ carrier G" "set fb \ carrier G" shows "wfactors G (a # fb) (a \ b)" using assms unfolding wfactors_def by (simp add: mult_cong_r) lemma (in monoid) factors_mult: assumes factors: "factors G fa a" "factors G fb b" and ascarr: "set fa \ carrier G" and bscarr: "set fb \ carrier G" shows "factors G (fa @ fb) (a \ b)" proof - have "foldr (\) (fa @ fb) \ = foldr (\) fa \ \ foldr (\) fb \" if "set fa \ carrier G" "Ball (set fa) (irreducible G)" using that bscarr by (induct fa) (simp_all add: m_assoc) then show ?thesis using assms unfolding factors_def by force qed lemma (in comm_monoid_cancel) wfactors_mult [intro]: assumes asf: "wfactors G as a" and bsf:"wfactors G bs b" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ascarr: "set as \ carrier G" and bscarr:"set bs \ carrier G" shows "wfactors G (as @ bs) (a \ b)" using wfactors_factors[OF asf ascarr] and wfactors_factors[OF bsf bscarr] proof clarsimp fix a' b' assume asf': "factors G as a'" and a'a: "a' \ a" and bsf': "factors G bs b'" and b'b: "b' \ b" from asf' have a'carr: "a' \ carrier G" by (rule factors_closed) fact from bsf' have b'carr: "b' \ carrier G" by (rule factors_closed) fact note carr = acarr bcarr a'carr b'carr ascarr bscarr from asf' bsf' have "factors G (as @ bs) (a' \ b')" by (rule factors_mult) fact+ with carr have abf': "wfactors G (as @ bs) (a' \ b')" by (intro factors_wfactors) simp_all also from b'b carr have trb: "a' \ b' \ a' \ b" by (intro mult_cong_r) also from a'a carr have tra: "a' \ b \ a \ b" by (intro mult_cong_l) finally show "wfactors G (as @ bs) (a \ b)" by (simp add: carr) qed lemma (in comm_monoid) factors_dividesI: assumes "factors G fs a" and "f \ set fs" and "set fs \ carrier G" shows "f divides a" using assms by (fast elim: factorsE intro: multlist_dividesI) lemma (in comm_monoid) wfactors_dividesI: assumes p: "wfactors G fs a" and fscarr: "set fs \ carrier G" and acarr: "a \ carrier G" and f: "f \ set fs" shows "f divides a" using wfactors_factors[OF p fscarr] proof clarsimp fix a' assume fsa': "factors G fs a'" and a'a: "a' \ a" with fscarr have a'carr: "a' \ carrier G" by (simp add: factors_closed) from fsa' fscarr f have "f divides a'" by (fast intro: factors_dividesI) also note a'a finally show "f divides a" by (simp add: f fscarr[THEN subsetD] acarr a'carr) qed subsubsection \Factorial monoids and wfactors\ lemma (in comm_monoid_cancel) factorial_monoidI: assumes wfactors_exists: "\a. \ a \ carrier G; a \ Units G \ \ \fs. set fs \ carrier G \ wfactors G fs a" and wfactors_unique: "\a fs fs'. \a \ carrier G; set fs \ carrier G; set fs' \ carrier G; wfactors G fs a; wfactors G fs' a\ \ essentially_equal G fs fs'" shows "factorial_monoid G" proof fix a assume acarr: "a \ carrier G" and anunit: "a \ Units G" from wfactors_exists[OF acarr anunit] obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast from wfactors_factors [OF afs ascarr] obtain a' where afs': "factors G as a'" and a'a: "a' \ a" by blast from afs' ascarr have a'carr: "a' \ carrier G" by fast have a'nunit: "a' \ Units G" proof clarify assume "a' \ Units G" also note a'a finally have "a \ Units G" by (simp add: acarr) with anunit show False .. qed from a'carr acarr a'a obtain u where uunit: "u \ Units G" and a': "a' = a \ u" by (blast elim: associatedE2) note [simp] = acarr Units_closed[OF uunit] Units_inv_closed[OF uunit] have "a = a \ \" by simp also have "\ = a \ (u \ inv u)" by (simp add: uunit) also have "\ = a' \ inv u" by (simp add: m_assoc[symmetric] a'[symmetric]) finally have a: "a = a' \ inv u" . from ascarr uunit have cr: "set (as[0:=(as!0 \ inv u)]) \ carrier G" by (cases as) auto from afs' uunit a'nunit acarr ascarr have "factors G (as[0:=(as!0 \ inv u)]) a" by (simp add: a factors_cong_unit) with cr show "\fs. set fs \ carrier G \ factors G fs a" by fast qed (blast intro: factors_wfactors wfactors_unique) subsection \Factorizations as Multisets\ text \Gives useful operations like intersection\ (* FIXME: use class_of x instead of closure_of {x} *) abbreviation "assocs G x \ eq_closure_of (division_rel G) {x}" definition "fmset G as = mset (map (assocs G) as)" text \Helper lemmas\ lemma (in monoid) assocs_repr_independence: assumes "y \ assocs G x" "x \ carrier G" shows "assocs G x = assocs G y" using assms by (simp add: eq_closure_of_def elem_def) (use associated_sym associated_trans in \blast+\) lemma (in monoid) assocs_self: assumes "x \ carrier G" shows "x \ assocs G x" using assms by (fastforce intro: closure_ofI2) lemma (in monoid) assocs_repr_independenceD: assumes repr: "assocs G x = assocs G y" and ycarr: "y \ carrier G" shows "y \ assocs G x" unfolding repr using ycarr by (intro assocs_self) lemma (in comm_monoid) assocs_assoc: assumes "a \ assocs G b" "b \ carrier G" shows "a \ b" using assms by (elim closure_ofE2) simp lemmas (in comm_monoid) assocs_eqD = assocs_repr_independenceD[THEN assocs_assoc] subsubsection \Comparing multisets\ lemma (in monoid) fmset_perm_cong: assumes prm: "as <~~> bs" shows "fmset G as = fmset G bs" - using perm_map[OF prm] unfolding mset_eq_perm fmset_def by blast + using perm_map[OF prm] unfolding fmset_def by blast lemma (in comm_monoid_cancel) eqc_listassoc_cong: assumes "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "map (assocs G) as = map (assocs G) bs" using assms proof (induction as arbitrary: bs) case Nil then show ?case by simp next case (Cons a as) then show ?case proof (clarsimp simp add: Cons_eq_map_conv list_all2_Cons1) fix z zs assume zzs: "a \ carrier G" "set as \ carrier G" "bs = z # zs" "a \ z" "as [\] zs" "z \ carrier G" "set zs \ carrier G" then show "assocs G a = assocs G z" apply (simp add: eq_closure_of_def elem_def) using \a \ carrier G\ \z \ carrier G\ \a \ z\ associated_sym associated_trans by blast+ qed qed lemma (in comm_monoid_cancel) fmset_listassoc_cong: assumes "as [\] bs" and "set as \ carrier G" and "set bs \ carrier G" shows "fmset G as = fmset G bs" using assms unfolding fmset_def by (simp add: eqc_listassoc_cong) lemma (in comm_monoid_cancel) ee_fmset: assumes ee: "essentially_equal G as bs" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "fmset G as = fmset G bs" using ee thm essentially_equal_def proof (elim essentially_equalE) fix as' assume prm: "as <~~> as'" from prm ascarr have as'carr: "set as' \ carrier G" by (rule perm_closed) from prm have "fmset G as = fmset G as'" by (rule fmset_perm_cong) also assume "as' [\] bs" with as'carr bscarr have "fmset G as' = fmset G bs" by (simp add: fmset_listassoc_cong) finally show "fmset G as = fmset G bs" . qed lemma (in comm_monoid_cancel) fmset_ee: assumes mset: "fmset G as = fmset G bs" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "essentially_equal G as bs" proof - from mset have "mset (map (assocs G) bs) = mset (map (assocs G) as)" by (simp add: fmset_def) then obtain p where \p permutes {.. \permute_list p (map (assocs G) as) = map (assocs G) bs\ by (rule mset_eq_permutation) then have \p permutes {.. \map (assocs G) (permute_list p as) = map (assocs G) bs\ by (simp_all add: permute_list_map) moreover define as' where \as' = permute_list p as\ ultimately have tp: "as <~~> as'" and tm: "map (assocs G) as' = map (assocs G) bs" - by (simp_all add: perm_iff_eq_mset) + by simp_all from tp show ?thesis proof (rule essentially_equalI) from tm tp ascarr have as'carr: "set as' \ carrier G" using perm_closed by blast from tm as'carr[THEN subsetD] bscarr[THEN subsetD] show "as' [\] bs" by (induct as' arbitrary: bs) (simp, fastforce dest: assocs_eqD[THEN associated_sym]) qed qed lemma (in comm_monoid_cancel) ee_is_fmset: assumes "set as \ carrier G" and "set bs \ carrier G" shows "essentially_equal G as bs = (fmset G as = fmset G bs)" using assms by (fast intro: ee_fmset fmset_ee) subsubsection \Interpreting multisets as factorizations\ lemma (in monoid) mset_fmsetEx: assumes elems: "\X. X \ set_mset Cs \ \x. P x \ X = assocs G x" shows "\cs. (\c \ set cs. P c) \ fmset G cs = Cs" proof - from surjE[OF surj_mset] obtain Cs' where Cs: "Cs = mset Cs'" by blast have "\cs. (\c \ set cs. P c) \ mset (map (assocs G) cs) = Cs" using elems unfolding Cs proof (induction Cs') case (Cons a Cs') then obtain c cs where csP: "\x\set cs. P x" and mset: "mset (map (assocs G) cs) = mset Cs'" and cP: "P c" and a: "a = assocs G c" by force then have tP: "\x\set (c#cs). P x" by simp show ?case using tP mset a by fastforce qed auto then show ?thesis by (simp add: fmset_def) qed lemma (in monoid) mset_wfactorsEx: assumes elems: "\X. X \ set_mset Cs \ \x. (x \ carrier G \ irreducible G x) \ X = assocs G x" shows "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = Cs" proof - have "\cs. (\c\set cs. c \ carrier G \ irreducible G c) \ fmset G cs = Cs" by (intro mset_fmsetEx, rule elems) then obtain cs where p[rule_format]: "\c\set cs. c \ carrier G \ irreducible G c" and Cs[symmetric]: "fmset G cs = Cs" by auto from p have cscarr: "set cs \ carrier G" by fast from p have "\c. c \ carrier G \ wfactors G cs c" by (intro wfactors_prod_exists) auto then obtain c where ccarr: "c \ carrier G" and cfs: "wfactors G cs c" by auto with cscarr Cs show ?thesis by fast qed subsubsection \Multiplication on multisets\ lemma (in factorial_monoid) mult_wfactors_fmset: assumes afs: "wfactors G as a" and bfs: "wfactors G bs b" and cfs: "wfactors G cs (a \ b)" and carr: "a \ carrier G" "b \ carrier G" "set as \ carrier G" "set bs \ carrier G" "set cs \ carrier G" shows "fmset G cs = fmset G as + fmset G bs" proof - from assms have "wfactors G (as @ bs) (a \ b)" by (intro wfactors_mult) with carr cfs have "essentially_equal G cs (as@bs)" by (intro ee_wfactorsI[of "a\b" "a\b"]) simp_all with carr have "fmset G cs = fmset G (as@bs)" by (intro ee_fmset) simp_all also have "fmset G (as@bs) = fmset G as + fmset G bs" by (simp add: fmset_def) finally show "fmset G cs = fmset G as + fmset G bs" . qed lemma (in factorial_monoid) mult_factors_fmset: assumes afs: "factors G as a" and bfs: "factors G bs b" and cfs: "factors G cs (a \ b)" and "set as \ carrier G" "set bs \ carrier G" "set cs \ carrier G" shows "fmset G cs = fmset G as + fmset G bs" using assms by (blast intro: factors_wfactors mult_wfactors_fmset) lemma (in comm_monoid_cancel) fmset_wfactors_mult: assumes mset: "fmset G cs = fmset G as + fmset G bs" and carr: "a \ carrier G" "b \ carrier G" "c \ carrier G" "set as \ carrier G" "set bs \ carrier G" "set cs \ carrier G" and fs: "wfactors G as a" "wfactors G bs b" "wfactors G cs c" shows "c \ a \ b" proof - from carr fs have m: "wfactors G (as @ bs) (a \ b)" by (intro wfactors_mult) from mset have "fmset G cs = fmset G (as@bs)" by (simp add: fmset_def) then have "essentially_equal G cs (as@bs)" by (rule fmset_ee) (simp_all add: carr) then show "c \ a \ b" by (rule ee_wfactorsD[of "cs" "as@bs"]) (simp_all add: assms m) qed subsubsection \Divisibility on multisets\ lemma (in factorial_monoid) divides_fmsubset: assumes ab: "a divides b" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and carr: "a \ carrier G" "b \ carrier G" "set as \ carrier G" "set bs \ carrier G" shows "fmset G as \# fmset G bs" using ab proof (elim dividesE) fix c assume ccarr: "c \ carrier G" from wfactors_exist [OF this] obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" by blast note carr = carr ccarr cscarr assume "b = a \ c" with afs bfs cfs carr have "fmset G bs = fmset G as + fmset G cs" by (intro mult_wfactors_fmset[OF afs cfs]) simp_all then show ?thesis by simp qed lemma (in comm_monoid_cancel) fmsubset_divides: assumes msubset: "fmset G as \# fmset G bs" and afs: "wfactors G as a" and bfs: "wfactors G bs b" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ascarr: "set as \ carrier G" and bscarr: "set bs \ carrier G" shows "a divides b" proof - from afs have airr: "\a \ set as. irreducible G a" by (fast elim: wfactorsE) from bfs have birr: "\b \ set bs. irreducible G b" by (fast elim: wfactorsE) have "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = fmset G bs - fmset G as" proof (intro mset_wfactorsEx, simp) fix X assume "X \# fmset G bs - fmset G as" then have "X \# fmset G bs" by (rule in_diffD) then have "X \ set (map (assocs G) bs)" by (simp add: fmset_def) then have "\x. x \ set bs \ X = assocs G x" by (induct bs) auto then obtain x where xbs: "x \ set bs" and X: "X = assocs G x" by auto with bscarr have xcarr: "x \ carrier G" by fast from xbs birr have xirr: "irreducible G x" by simp from xcarr and xirr and X show "\x. x \ carrier G \ irreducible G x \ X = assocs G x" by fast qed then obtain c cs where ccarr: "c \ carrier G" and cscarr: "set cs \ carrier G" and csf: "wfactors G cs c" and csmset: "fmset G cs = fmset G bs - fmset G as" by auto from csmset msubset have "fmset G bs = fmset G as + fmset G cs" by (simp add: multiset_eq_iff subseteq_mset_def) then have basc: "b \ a \ c" by (rule fmset_wfactors_mult) fact+ then show ?thesis proof (elim associatedE2) fix u assume "u \ Units G" "b = a \ c \ u" with acarr ccarr show "a divides b" by (fast intro: dividesI[of "c \ u"] m_assoc) qed (simp_all add: acarr bcarr ccarr) qed lemma (in factorial_monoid) divides_as_fmsubset: assumes "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "a divides b = (fmset G as \# fmset G bs)" using assms by (blast intro: divides_fmsubset fmsubset_divides) text \Proper factors on multisets\ lemma (in factorial_monoid) fmset_properfactor: assumes asubb: "fmset G as \# fmset G bs" and anb: "fmset G as \ fmset G bs" and "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "properfactor G a b" proof (rule properfactorI) show "a divides b" using assms asubb fmsubset_divides by blast show "\ b divides a" by (meson anb assms asubb factorial_monoid.divides_fmsubset factorial_monoid_axioms subset_mset.antisym) qed lemma (in factorial_monoid) properfactor_fmset: assumes "properfactor G a b" and "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "fmset G as \# fmset G bs" using assms by (meson divides_as_fmsubset properfactor_divides) lemma (in factorial_monoid) properfactor_fmset_ne: assumes pf: "properfactor G a b" and "wfactors G as a" and "wfactors G bs b" and "a \ carrier G" and "b \ carrier G" and "set as \ carrier G" and "set bs \ carrier G" shows "fmset G as \ fmset G bs" using properfactorE [OF pf] assms divides_as_fmsubset by force subsection \Irreducible Elements are Prime\ lemma (in factorial_monoid) irreducible_prime: assumes pirr: "irreducible G p" and pcarr: "p \ carrier G" shows "prime G p" using pirr proof (elim irreducibleE, intro primeI) fix a b assume acarr: "a \ carrier G" and bcarr: "b \ carrier G" and pdvdab: "p divides (a \ b)" and pnunit: "p \ Units G" assume irreduc[rule_format]: "\b. b \ carrier G \ properfactor G b p \ b \ Units G" from pdvdab obtain c where ccarr: "c \ carrier G" and abpc: "a \ b = p \ c" by (rule dividesE) obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" using wfactors_exist [OF acarr] by blast obtain bs where bscarr: "set bs \ carrier G" and bfs: "wfactors G bs b" using wfactors_exist [OF bcarr] by blast obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" using wfactors_exist [OF ccarr] by blast note carr[simp] = pcarr acarr bcarr ccarr ascarr bscarr cscarr from pirr cfs abpc have "wfactors G (p # cs) (a \ b)" by (simp add: wfactors_mult_single) moreover have "wfactors G (as @ bs) (a \ b)" by (rule wfactors_mult [OF afs bfs]) fact+ ultimately have "essentially_equal G (p # cs) (as @ bs)" by (rule wfactors_unique) simp+ then obtain ds where "p # cs <~~> ds" and dsassoc: "ds [\] (as @ bs)" by (fast elim: essentially_equalE) then have "p \ set ds" - by (simp add: perm_set_eq[symmetric]) + by (metis \mset (p # cs) = mset ds\ insert_iff list.set(2) perm_set_eq) with dsassoc obtain p' where "p' \ set (as@bs)" and pp': "p \ p'" unfolding list_all2_conv_all_nth set_conv_nth by force then consider "p' \ set as" | "p' \ set bs" by auto then show "p divides a \ p divides b" using afs bfs divides_cong_l pp' wfactors_dividesI by (meson acarr ascarr bcarr bscarr pcarr) qed \ \A version using \<^const>\factors\, more complicated\ lemma (in factorial_monoid) factors_irreducible_prime: assumes pirr: "irreducible G p" and pcarr: "p \ carrier G" shows "prime G p" proof (rule primeI) show "p \ Units G" by (meson irreducibleE pirr) have irreduc: "\b. \b \ carrier G; properfactor G b p\ \ b \ Units G" using pirr by (auto simp: irreducible_def) show "p divides a \ p divides b" if acarr: "a \ carrier G" and bcarr: "b \ carrier G" and pdvdab: "p divides (a \ b)" for a b proof - from pdvdab obtain c where ccarr: "c \ carrier G" and abpc: "a \ b = p \ c" by (rule dividesE) note [simp] = pcarr acarr bcarr ccarr show "p divides a \ p divides b" proof (cases "a \ Units G") case True then have "p divides b" by (metis acarr associatedI2' associated_def bcarr divides_trans m_comm pcarr pdvdab) then show ?thesis .. next case anunit: False show ?thesis proof (cases "b \ Units G") case True then have "p divides a" by (meson acarr bcarr divides_unit irreducible_prime pcarr pdvdab pirr prime_def) then show ?thesis .. next case bnunit: False then have cnunit: "c \ Units G" by (metis abpc acarr anunit bcarr ccarr irreducible_prodE irreducible_prod_rI pcarr pirr) then have abnunit: "a \ b \ Units G" using acarr anunit bcarr unit_factor by blast obtain as where ascarr: "set as \ carrier G" and afac: "factors G as a" using factors_exist [OF acarr anunit] by blast obtain bs where bscarr: "set bs \ carrier G" and bfac: "factors G bs b" using factors_exist [OF bcarr bnunit] by blast obtain cs where cscarr: "set cs \ carrier G" and cfac: "factors G cs c" using factors_exist [OF ccarr cnunit] by auto note [simp] = ascarr bscarr cscarr from pirr cfac abpc have abfac': "factors G (p # cs) (a \ b)" by (simp add: factors_mult_single) from afac and bfac have "factors G (as @ bs) (a \ b)" by (rule factors_mult) fact+ with abfac' have "essentially_equal G (p # cs) (as @ bs)" using abnunit factors_unique by auto then obtain ds where "p # cs <~~> ds" and dsassoc: "ds [\] (as @ bs)" by (fast elim: essentially_equalE) then have "p \ set ds" - by (simp add: perm_set_eq[symmetric]) + by (metis list.set_intros(1) set_mset_mset) with dsassoc obtain p' where "p' \ set (as@bs)" and pp': "p \ p'" unfolding list_all2_conv_all_nth set_conv_nth by force then consider "p' \ set as" | "p' \ set bs" by auto then show "p divides a \ p divides b" by (meson afac bfac divides_cong_l factors_dividesI pp' ascarr bscarr pcarr) qed qed qed qed subsection \Greatest Common Divisors and Lowest Common Multiples\ subsubsection \Definitions\ definition isgcd :: "[('a,_) monoid_scheme, 'a, 'a, 'a] \ bool" ("(_ gcdof\ _ _)" [81,81,81] 80) where "x gcdof\<^bsub>G\<^esub> a b \ x divides\<^bsub>G\<^esub> a \ x divides\<^bsub>G\<^esub> b \ (\y\carrier G. (y divides\<^bsub>G\<^esub> a \ y divides\<^bsub>G\<^esub> b \ y divides\<^bsub>G\<^esub> x))" definition islcm :: "[_, 'a, 'a, 'a] \ bool" ("(_ lcmof\ _ _)" [81,81,81] 80) where "x lcmof\<^bsub>G\<^esub> a b \ a divides\<^bsub>G\<^esub> x \ b divides\<^bsub>G\<^esub> x \ (\y\carrier G. (a divides\<^bsub>G\<^esub> y \ b divides\<^bsub>G\<^esub> y \ x divides\<^bsub>G\<^esub> y))" definition somegcd :: "('a,_) monoid_scheme \ 'a \ 'a \ 'a" where "somegcd G a b = (SOME x. x \ carrier G \ x gcdof\<^bsub>G\<^esub> a b)" definition somelcm :: "('a,_) monoid_scheme \ 'a \ 'a \ 'a" where "somelcm G a b = (SOME x. x \ carrier G \ x lcmof\<^bsub>G\<^esub> a b)" definition "SomeGcd G A = Lattice.inf (division_rel G) A" locale gcd_condition_monoid = comm_monoid_cancel + assumes gcdof_exists: "\a \ carrier G; b \ carrier G\ \ \c. c \ carrier G \ c gcdof a b" locale primeness_condition_monoid = comm_monoid_cancel + assumes irreducible_prime: "\a \ carrier G; irreducible G a\ \ prime G a" locale divisor_chain_condition_monoid = comm_monoid_cancel + assumes division_wellfounded: "wf {(x, y). x \ carrier G \ y \ carrier G \ properfactor G x y}" subsubsection \Connections to \texttt{Lattice.thy}\ lemma gcdof_greatestLower: fixes G (structure) assumes carr[simp]: "a \ carrier G" "b \ carrier G" shows "(x \ carrier G \ x gcdof a b) = greatest (division_rel G) x (Lower (division_rel G) {a, b})" by (auto simp: isgcd_def greatest_def Lower_def elem_def) lemma lcmof_leastUpper: fixes G (structure) assumes carr[simp]: "a \ carrier G" "b \ carrier G" shows "(x \ carrier G \ x lcmof a b) = least (division_rel G) x (Upper (division_rel G) {a, b})" by (auto simp: islcm_def least_def Upper_def elem_def) lemma somegcd_meet: fixes G (structure) assumes carr: "a \ carrier G" "b \ carrier G" shows "somegcd G a b = meet (division_rel G) a b" by (simp add: somegcd_def meet_def inf_def gcdof_greatestLower[OF carr]) lemma (in monoid) isgcd_divides_l: assumes "a divides b" and "a \ carrier G" "b \ carrier G" shows "a gcdof a b" using assms unfolding isgcd_def by fast lemma (in monoid) isgcd_divides_r: assumes "b divides a" and "a \ carrier G" "b \ carrier G" shows "b gcdof a b" using assms unfolding isgcd_def by fast subsubsection \Existence of gcd and lcm\ lemma (in factorial_monoid) gcdof_exists: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" shows "\c. c \ carrier G \ c gcdof a b" proof - from wfactors_exist [OF acarr] obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast from afs have airr: "\a \ set as. irreducible G a" by (fast elim: wfactorsE) from wfactors_exist [OF bcarr] obtain bs where bscarr: "set bs \ carrier G" and bfs: "wfactors G bs b" by blast from bfs have birr: "\b \ set bs. irreducible G b" by (fast elim: wfactorsE) have "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = fmset G as \# fmset G bs" proof (intro mset_wfactorsEx) fix X assume "X \# fmset G as \# fmset G bs" then have "X \# fmset G as" by simp then have "X \ set (map (assocs G) as)" by (simp add: fmset_def) then have "\x. X = assocs G x \ x \ set as" by (induct as) auto then obtain x where X: "X = assocs G x" and xas: "x \ set as" by blast with ascarr have xcarr: "x \ carrier G" by blast from xas airr have xirr: "irreducible G x" by simp from xcarr and xirr and X show "\x. (x \ carrier G \ irreducible G x) \ X = assocs G x" by blast qed then obtain c cs where ccarr: "c \ carrier G" and cscarr: "set cs \ carrier G" and csirr: "wfactors G cs c" and csmset: "fmset G cs = fmset G as \# fmset G bs" by auto have "c gcdof a b" proof (simp add: isgcd_def, safe) from csmset have "fmset G cs \# fmset G as" by simp then show "c divides a" by (rule fmsubset_divides) fact+ next from csmset have "fmset G cs \# fmset G bs" by simp then show "c divides b" by (rule fmsubset_divides) fact+ next fix y assume "y \ carrier G" from wfactors_exist [OF this] obtain ys where yscarr: "set ys \ carrier G" and yfs: "wfactors G ys y" by blast assume "y divides a" then have ya: "fmset G ys \# fmset G as" by (rule divides_fmsubset) fact+ assume "y divides b" then have yb: "fmset G ys \# fmset G bs" by (rule divides_fmsubset) fact+ from ya yb csmset have "fmset G ys \# fmset G cs" by (simp add: subset_mset_def) then show "y divides c" by (rule fmsubset_divides) fact+ qed with ccarr show "\c. c \ carrier G \ c gcdof a b" by fast qed lemma (in factorial_monoid) lcmof_exists: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" shows "\c. c \ carrier G \ c lcmof a b" proof - from wfactors_exist [OF acarr] obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast from afs have airr: "\a \ set as. irreducible G a" by (fast elim: wfactorsE) from wfactors_exist [OF bcarr] obtain bs where bscarr: "set bs \ carrier G" and bfs: "wfactors G bs b" by blast from bfs have birr: "\b \ set bs. irreducible G b" by (fast elim: wfactorsE) have "\c cs. c \ carrier G \ set cs \ carrier G \ wfactors G cs c \ fmset G cs = (fmset G as - fmset G bs) + fmset G bs" proof (intro mset_wfactorsEx) fix X assume "X \# (fmset G as - fmset G bs) + fmset G bs" then have "X \# fmset G as \ X \# fmset G bs" by (auto dest: in_diffD) then consider "X \ set_mset (fmset G as)" | "X \ set_mset (fmset G bs)" by fast then show "\x. (x \ carrier G \ irreducible G x) \ X = assocs G x" proof cases case 1 then have "X \ set (map (assocs G) as)" by (simp add: fmset_def) then have "\x. x \ set as \ X = assocs G x" by (induct as) auto then obtain x where xas: "x \ set as" and X: "X = assocs G x" by auto with ascarr have xcarr: "x \ carrier G" by fast from xas airr have xirr: "irreducible G x" by simp from xcarr and xirr and X show ?thesis by fast next case 2 then have "X \ set (map (assocs G) bs)" by (simp add: fmset_def) then have "\x. x \ set bs \ X = assocs G x" by (induct as) auto then obtain x where xbs: "x \ set bs" and X: "X = assocs G x" by auto with bscarr have xcarr: "x \ carrier G" by fast from xbs birr have xirr: "irreducible G x" by simp from xcarr and xirr and X show ?thesis by fast qed qed then obtain c cs where ccarr: "c \ carrier G" and cscarr: "set cs \ carrier G" and csirr: "wfactors G cs c" and csmset: "fmset G cs = fmset G as - fmset G bs + fmset G bs" by auto have "c lcmof a b" proof (simp add: islcm_def, safe) from csmset have "fmset G as \# fmset G cs" by (simp add: subseteq_mset_def, force) then show "a divides c" by (rule fmsubset_divides) fact+ next from csmset have "fmset G bs \# fmset G cs" by (simp add: subset_mset_def) then show "b divides c" by (rule fmsubset_divides) fact+ next fix y assume "y \ carrier G" from wfactors_exist [OF this] obtain ys where yscarr: "set ys \ carrier G" and yfs: "wfactors G ys y" by blast assume "a divides y" then have ya: "fmset G as \# fmset G ys" by (rule divides_fmsubset) fact+ assume "b divides y" then have yb: "fmset G bs \# fmset G ys" by (rule divides_fmsubset) fact+ from ya yb csmset have "fmset G cs \# fmset G ys" using subset_eq_diff_conv subset_mset.le_diff_conv2 by fastforce then show "c divides y" by (rule fmsubset_divides) fact+ qed with ccarr show "\c. c \ carrier G \ c lcmof a b" by fast qed subsection \Conditions for Factoriality\ subsubsection \Gcd condition\ lemma (in gcd_condition_monoid) division_weak_lower_semilattice [simp]: "weak_lower_semilattice (division_rel G)" proof - interpret weak_partial_order "division_rel G" .. show ?thesis proof (unfold_locales, simp_all) fix x y assume carr: "x \ carrier G" "y \ carrier G" from gcdof_exists [OF this] obtain z where zcarr: "z \ carrier G" and isgcd: "z gcdof x y" by blast with carr have "greatest (division_rel G) z (Lower (division_rel G) {x, y})" by (subst gcdof_greatestLower[symmetric], simp+) then show "\z. greatest (division_rel G) z (Lower (division_rel G) {x, y})" by fast qed qed lemma (in gcd_condition_monoid) gcdof_cong_l: assumes "a' \ a" "a gcdof b c" "a' \ carrier G" and carr': "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "a' gcdof b c" proof - interpret weak_lower_semilattice "division_rel G" by simp have "is_glb (division_rel G) a' {b, c}" by (subst greatest_Lower_cong_l[of _ a]) (simp_all add: assms gcdof_greatestLower[symmetric]) then have "a' \ carrier G \ a' gcdof b c" by (simp add: gcdof_greatestLower carr') then show ?thesis .. qed lemma (in gcd_condition_monoid) gcd_closed [simp]: assumes "a \ carrier G" "b \ carrier G" shows "somegcd G a b \ carrier G" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis using assms meet_closed by (simp add: somegcd_meet) qed lemma (in gcd_condition_monoid) gcd_isgcd: assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) gcdof a b" proof - interpret weak_lower_semilattice "division_rel G" by simp from assms have "somegcd G a b \ carrier G \ (somegcd G a b) gcdof a b" by (simp add: gcdof_greatestLower inf_of_two_greatest meet_def somegcd_meet) then show "(somegcd G a b) gcdof a b" by simp qed lemma (in gcd_condition_monoid) gcd_exists: assumes "a \ carrier G" "b \ carrier G" shows "\x\carrier G. x = somegcd G a b" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis assms gcd_closed) qed lemma (in gcd_condition_monoid) gcd_divides_l: assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) divides a" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis assms gcd_isgcd isgcd_def) qed lemma (in gcd_condition_monoid) gcd_divides_r: assumes "a \ carrier G" "b \ carrier G" shows "(somegcd G a b) divides b" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis assms gcd_isgcd isgcd_def) qed lemma (in gcd_condition_monoid) gcd_divides: assumes "z divides x" "z divides y" and L: "x \ carrier G" "y \ carrier G" "z \ carrier G" shows "z divides (somegcd G x y)" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (metis gcd_isgcd isgcd_def assms) qed lemma (in gcd_condition_monoid) gcd_cong_l: assumes "x \ x'" "x \ carrier G" "x' \ carrier G" "y \ carrier G" shows "somegcd G x y \ somegcd G x' y" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis using somegcd_meet assms by (metis eq_object.select_convs(1) meet_cong_l partial_object.select_convs(1)) qed lemma (in gcd_condition_monoid) gcd_cong_r: assumes "y \ y'" "x \ carrier G" "y \ carrier G" "y' \ carrier G" shows "somegcd G x y \ somegcd G x y'" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis by (meson associated_def assms gcd_closed gcd_divides gcd_divides_l gcd_divides_r monoid.divides_trans monoid_axioms) qed lemma (in gcd_condition_monoid) gcdI: assumes dvd: "a divides b" "a divides c" and others: "\y. \y\carrier G; y divides b; y divides c\ \ y divides a" and acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ccarr: "c \ carrier G" shows "a \ somegcd G b c" proof - have "\a. a \ carrier G \ a gcdof b c" by (simp add: bcarr ccarr gcdof_exists) moreover have "\x. x \ carrier G \ x gcdof b c \ a \ x" by (simp add: acarr associated_def dvd isgcd_def others) ultimately show ?thesis unfolding somegcd_def by (blast intro: someI2_ex) qed lemma (in gcd_condition_monoid) gcdI2: assumes "a gcdof b c" and "a \ carrier G" and "b \ carrier G" and "c \ carrier G" shows "a \ somegcd G b c" using assms unfolding isgcd_def by (simp add: gcdI) lemma (in gcd_condition_monoid) SomeGcd_ex: assumes "finite A" "A \ carrier G" "A \ {}" shows "\x \ carrier G. x = SomeGcd G A" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis using finite_inf_closed by (simp add: assms SomeGcd_def) qed lemma (in gcd_condition_monoid) gcd_assoc: assumes "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "somegcd G (somegcd G a b) c \ somegcd G a (somegcd G b c)" proof - interpret weak_lower_semilattice "division_rel G" by simp show ?thesis unfolding associated_def by (meson assms divides_trans gcd_divides gcd_divides_l gcd_divides_r gcd_exists) qed lemma (in gcd_condition_monoid) gcd_mult: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" and ccarr: "c \ carrier G" shows "c \ somegcd G a b \ somegcd G (c \ a) (c \ b)" proof - (* following Jacobson, Basic Algebra, p.140 *) let ?d = "somegcd G a b" let ?e = "somegcd G (c \ a) (c \ b)" note carr[simp] = acarr bcarr ccarr have dcarr: "?d \ carrier G" by simp have ecarr: "?e \ carrier G" by simp note carr = carr dcarr ecarr have "?d divides a" by (simp add: gcd_divides_l) then have cd'ca: "c \ ?d divides (c \ a)" by (simp add: divides_mult_lI) have "?d divides b" by (simp add: gcd_divides_r) then have cd'cb: "c \ ?d divides (c \ b)" by (simp add: divides_mult_lI) from cd'ca cd'cb have cd'e: "c \ ?d divides ?e" by (rule gcd_divides) simp_all then obtain u where ucarr[simp]: "u \ carrier G" and e_cdu: "?e = c \ ?d \ u" by blast note carr = carr ucarr have "?e divides c \ a" by (rule gcd_divides_l) simp_all then obtain x where xcarr: "x \ carrier G" and ca_ex: "c \ a = ?e \ x" by blast with e_cdu have ca_cdux: "c \ a = c \ ?d \ u \ x" by simp from ca_cdux xcarr have "c \ a = c \ (?d \ u \ x)" by (simp add: m_assoc) then have "a = ?d \ u \ x" by (rule l_cancel[of c a]) (simp add: xcarr)+ then have du'a: "?d \ u divides a" by (rule dividesI[OF xcarr]) have "?e divides c \ b" by (intro gcd_divides_r) simp_all then obtain x where xcarr: "x \ carrier G" and cb_ex: "c \ b = ?e \ x" by blast with e_cdu have cb_cdux: "c \ b = c \ ?d \ u \ x" by simp from cb_cdux xcarr have "c \ b = c \ (?d \ u \ x)" by (simp add: m_assoc) with xcarr have "b = ?d \ u \ x" by (intro l_cancel[of c b]) simp_all then have du'b: "?d \ u divides b" by (intro dividesI[OF xcarr]) from du'a du'b carr have du'd: "?d \ u divides ?d" by (intro gcd_divides) simp_all then have uunit: "u \ Units G" proof (elim dividesE) fix v assume vcarr[simp]: "v \ carrier G" assume d: "?d = ?d \ u \ v" have "?d \ \ = ?d \ u \ v" by simp fact also have "?d \ u \ v = ?d \ (u \ v)" by (simp add: m_assoc) finally have "?d \ \ = ?d \ (u \ v)" . then have i2: "\ = u \ v" by (rule l_cancel) simp_all then have i1: "\ = v \ u" by (simp add: m_comm) from vcarr i1[symmetric] i2[symmetric] show "u \ Units G" by (auto simp: Units_def) qed from e_cdu uunit have "somegcd G (c \ a) (c \ b) \ c \ somegcd G a b" by (intro associatedI2[of u]) simp_all from this[symmetric] show "c \ somegcd G a b \ somegcd G (c \ a) (c \ b)" by simp qed lemma (in monoid) assoc_subst: assumes ab: "a \ b" and cP: "\a b. a \ carrier G \ b \ carrier G \ a \ b \ f a \ carrier G \ f b \ carrier G \ f a \ f b" and carr: "a \ carrier G" "b \ carrier G" shows "f a \ f b" using assms by auto lemma (in gcd_condition_monoid) relprime_mult: assumes abrelprime: "somegcd G a b \ \" and acrelprime: "somegcd G a c \ \" and carr[simp]: "a \ carrier G" "b \ carrier G" "c \ carrier G" shows "somegcd G a (b \ c) \ \" proof - have "c = c \ \" by simp also from abrelprime[symmetric] have "\ \ c \ somegcd G a b" by (rule assoc_subst) (simp add: mult_cong_r)+ also have "\ \ somegcd G (c \ a) (c \ b)" by (rule gcd_mult) fact+ finally have c: "c \ somegcd G (c \ a) (c \ b)" by simp from carr have a: "a \ somegcd G a (c \ a)" by (fast intro: gcdI divides_prod_l) have "somegcd G a (b \ c) \ somegcd G a (c \ b)" by (simp add: m_comm) also from a have "\ \ somegcd G (somegcd G a (c \ a)) (c \ b)" by (rule assoc_subst) (simp add: gcd_cong_l)+ also from gcd_assoc have "\ \ somegcd G a (somegcd G (c \ a) (c \ b))" by (rule assoc_subst) simp+ also from c[symmetric] have "\ \ somegcd G a c" by (rule assoc_subst) (simp add: gcd_cong_r)+ also note acrelprime finally show "somegcd G a (b \ c) \ \" by simp qed lemma (in gcd_condition_monoid) primeness_condition: "primeness_condition_monoid G" proof - have *: "p divides a \ p divides b" if pcarr[simp]: "p \ carrier G" and acarr[simp]: "a \ carrier G" and bcarr[simp]: "b \ carrier G" and pirr: "irreducible G p" and pdvdab: "p divides a \ b" for p a b proof - from pirr have pnunit: "p \ Units G" and r: "\b. \b \ carrier G; properfactor G b p\ \ b \ Units G" by (fast elim: irreducibleE)+ show "p divides a \ p divides b" proof (rule ccontr, clarsimp) assume npdvda: "\ p divides a" and npdvdb: "\ p divides b" have "\ \ somegcd G p a" proof (intro gcdI unit_divides) show "\y. \y \ carrier G; y divides p; y divides a\ \ y \ Units G" by (meson divides_trans npdvda pcarr properfactorI r) qed auto with pcarr acarr have pa: "somegcd G p a \ \" by (fast intro: associated_sym[of "\"] gcd_closed) have "\ \ somegcd G p b" proof (intro gcdI unit_divides) show "\y. \y \ carrier G; y divides p; y divides b\ \ y \ Units G" by (meson divides_trans npdvdb pcarr properfactorI r) qed auto with pcarr bcarr have pb: "somegcd G p b \ \" by (fast intro: associated_sym[of "\"] gcd_closed) have "p \ somegcd G p (a \ b)" using pdvdab by (simp add: gcdI2 isgcd_divides_l) also from pa pb pcarr acarr bcarr have "somegcd G p (a \ b) \ \" by (rule relprime_mult) finally have "p \ \" by simp with pcarr have "p \ Units G" by (fast intro: assoc_unit_l) with pnunit show False .. qed qed show ?thesis by unfold_locales (metis * primeI irreducibleE) qed sublocale gcd_condition_monoid \ primeness_condition_monoid by (rule primeness_condition) subsubsection \Divisor chain condition\ lemma (in divisor_chain_condition_monoid) wfactors_exist: assumes acarr: "a \ carrier G" shows "\as. set as \ carrier G \ wfactors G as a" proof - have r: "a \ carrier G \ (\as. set as \ carrier G \ wfactors G as a)" using division_wellfounded proof (induction rule: wf_induct_rule) case (less x) then have xcarr: "x \ carrier G" by auto show ?case proof (cases "x \ Units G") case True then show ?thesis by (metis bot.extremum list.set(1) unit_wfactors) next case xnunit: False show ?thesis proof (cases "irreducible G x") case True then show ?thesis by (rule_tac x="[x]" in exI) (simp add: wfactors_def xcarr) next case False then obtain y where ycarr: "y \ carrier G" and ynunit: "y \ Units G" and pfyx: "properfactor G y x" by (meson irreducible_def xnunit) obtain ys where yscarr: "set ys \ carrier G" and yfs: "wfactors G ys y" using less ycarr pfyx by blast then obtain z where zcarr: "z \ carrier G" and x: "x = y \ z" by (meson dividesE pfyx properfactorE2) from zcarr ycarr have "properfactor G z x" using m_comm properfactorI3 x ynunit by blast with less zcarr obtain zs where zscarr: "set zs \ carrier G" and zfs: "wfactors G zs z" by blast from yscarr zscarr have xscarr: "set (ys@zs) \ carrier G" by simp have "wfactors G (ys@zs) (y\z)" using xscarr ycarr yfs zcarr zfs by auto then have "wfactors G (ys@zs) x" by (simp add: x) with xscarr show "\xs. set xs \ carrier G \ wfactors G xs x" by fast qed qed qed from acarr show ?thesis by (rule r) qed subsubsection \Primeness condition\ lemma (in comm_monoid_cancel) multlist_prime_pos: assumes aprime: "prime G a" and carr: "a \ carrier G" and as: "set as \ carrier G" "a divides (foldr (\) as \)" shows "\i carrier G" "set as \ carrier G" and "a divides x \ foldr (\) as \" by (auto simp: ) with carr aprime have "a divides x \ a divides foldr (\) as \" by (intro prime_divides) simp+ then show ?case using Cons.IH Cons.prems(1) by force qed proposition (in primeness_condition_monoid) wfactors_unique: assumes "wfactors G as a" "wfactors G as' a" and "a \ carrier G" "set as \ carrier G" "set as' \ carrier G" shows "essentially_equal G as as'" using assms proof (induct as arbitrary: a as') case Nil then have "a \ \" - by (meson Units_one_closed one_closed perm.Nil perm_wfactorsD unit_wfactors) + by (simp add: perm_wfactorsD) then have "as' = []" using Nil.prems assoc_unit_l unit_wfactors_empty by blast then show ?case by auto next case (Cons ah as) then have ahdvda: "ah divides a" using wfactors_dividesI by auto then obtain a' where a'carr: "a' \ carrier G" and a: "a = ah \ a'" by blast have carr_ah: "ah \ carrier G" "set as \ carrier G" using Cons.prems by fastforce+ have "ah \ foldr (\) as \ \ a" by (rule wfactorsE[OF \wfactors G (ah # as) a\]) auto then have "foldr (\) as \ \ a'" by (metis Cons.prems(4) a a'carr assoc_l_cancel insert_subset list.set(2) monoid.multlist_closed monoid_axioms) then have a'fs: "wfactors G as a'" by (meson Cons.prems(1) set_subset_Cons subset_iff wfactorsE wfactorsI) then have ahirr: "irreducible G ah" by (meson Cons.prems(1) list.set_intros(1) wfactorsE) with Cons have ahprime: "prime G ah" by (simp add: irreducible_prime) note ahdvda also have "a divides (foldr (\) as' \)" by (meson Cons.prems(2) associatedE wfactorsE) finally have "ah divides (foldr (\) as' \)" using Cons.prems(4) by auto with ahprime have "\i carrier G" and asi: "as'!i = ah \ x" by blast have irrasi: "irreducible G (as'!i)" using nth_mem[OF len] wfactorsE by (metis Cons.prems(2)) have asicarr[simp]: "as'!i \ carrier G" using len \set as' \ carrier G\ nth_mem by blast have asiah: "as'!i \ ah" by (metis \ah \ carrier G\ \x \ carrier G\ asi irrasi ahprime associatedI2 irreducible_prodE primeE) note setparts = set_take_subset[of i as'] set_drop_subset[of "Suc i" as'] have "\aa_1. aa_1 \ carrier G \ wfactors G (take i as') aa_1" using Cons by (metis setparts(1) subset_trans in_set_takeD wfactorsE wfactors_prod_exists) then obtain aa_1 where aa1carr [simp]: "aa_1 \ carrier G" and aa1fs: "wfactors G (take i as') aa_1" by auto obtain aa_2 where aa2carr [simp]: "aa_2 \ carrier G" and aa2fs: "wfactors G (drop (Suc i) as') aa_2" by (metis Cons.prems(2) Cons.prems(5) subset_code(1) in_set_dropD wfactors_def wfactors_prod_exists) have set_drop: "set (drop (Suc i) as') \ carrier G" using Cons.prems(5) setparts(2) by blast moreover have set_take: "set (take i as') \ carrier G" using Cons.prems(5) setparts by auto moreover have v1: "wfactors G (take i as' @ drop (Suc i) as') (aa_1 \ aa_2)" using aa1fs aa2fs \set as' \ carrier G\ by (force simp add: dest: in_set_takeD in_set_dropD) ultimately have v1': "wfactors G (as'!i # take i as' @ drop (Suc i) as') (as'!i \ (aa_1 \ aa_2))" using irrasi wfactors_mult_single by (simp add: irrasi v1 wfactors_mult_single) have "wfactors G (as'!i # drop (Suc i) as') (as'!i \ aa_2)" by (simp add: aa2fs irrasi set_drop wfactors_mult_single) with len aa1carr aa2carr aa1fs have v2: "wfactors G (take i as' @ as'!i # drop (Suc i) as') (aa_1 \ (as'!i \ aa_2))" using wfactors_mult by (simp add: set_take set_drop) from len have as': "as' = (take i as' @ as'!i # drop (Suc i) as')" by (simp add: Cons_nth_drop_Suc) have eer: "essentially_equal G (take i as' @ as'!i # drop (Suc i) as') as'" using Cons.prems(5) as' by auto with v2 aa1carr aa2carr nth_mem[OF len] have "aa_1 \ (as'!i \ aa_2) \ a" using Cons.prems as' comm_monoid_cancel.ee_wfactorsD is_comm_monoid_cancel by fastforce then have t1: "as'!i \ (aa_1 \ aa_2) \ a" by (metis aa1carr aa2carr asicarr m_lcomm) from asiah have "ah \ (aa_1 \ aa_2) \ as'!i \ (aa_1 \ aa_2)" by (simp add: \ah \ carrier G\ associated_sym mult_cong_l) also note t1 finally have "ah \ (aa_1 \ aa_2) \ a" using Cons.prems(3) carr_ah aa1carr aa2carr by blast with aa1carr aa2carr a'carr nth_mem[OF len] have a': "aa_1 \ aa_2 \ a'" using a assoc_l_cancel carr_ah(1) by blast note v1 also note a' finally have "wfactors G (take i as' @ drop (Suc i) as') a'" by (simp add: a'carr set_drop set_take) from a'fs this have "essentially_equal G as (take i as' @ drop (Suc i) as')" using Cons.hyps a'carr carr_ah(2) set_drop set_take by auto - with carr_ah have ee1: "essentially_equal G (ah # as) (ah # take i as' @ drop (Suc i) as')" - by (auto simp: essentially_equal_def) + then obtain bs where \mset as = mset bs\ \bs [\] take i as' @ drop (Suc i) as'\ + by (auto simp add: essentially_equal_def) + with carr_ah have \mset (ah # as) = mset (ah # bs)\ \ah # bs [\] ah # take i as' @ drop (Suc i) as'\ + by simp_all + then have ee1: "essentially_equal G (ah # as) (ah # take i as' @ drop (Suc i) as')" + unfolding essentially_equal_def by blast have ee2: "essentially_equal G (ah # take i as' @ drop (Suc i) as') (as' ! i # take i as' @ drop (Suc i) as')" proof (intro essentially_equalI) show "ah # take i as' @ drop (Suc i) as' <~~> ah # take i as' @ drop (Suc i) as'" by simp next show "ah # take i as' @ drop (Suc i) as' [\] as' ! i # take i as' @ drop (Suc i) as'" by (simp add: asiah associated_sym set_drop set_take) qed note ee1 also note ee2 also have "essentially_equal G (as' ! i # take i as' @ drop (Suc i) as') (take i as' @ as' ! i # drop (Suc i) as')" by (metis Cons.prems(5) as' essentially_equalI listassoc_refl perm_append_Cons) finally have "essentially_equal G (ah # as) (take i as' @ as' ! i # drop (Suc i) as')" using Cons.prems(4) set_drop set_take by auto then show ?case using as' by auto qed subsubsection \Application to factorial monoids\ text \Number of factors for wellfoundedness\ definition factorcount :: "_ \ 'a \ nat" where "factorcount G a = (THE c. \as. set as \ carrier G \ wfactors G as a \ c = length as)" lemma (in monoid) ee_length: assumes ee: "essentially_equal G as bs" shows "length as = length bs" by (rule essentially_equalE[OF ee]) (metis list_all2_conv_all_nth perm_length) lemma (in factorial_monoid) factorcount_exists: assumes carr[simp]: "a \ carrier G" shows "\c. \as. set as \ carrier G \ wfactors G as a \ c = length as" proof - have "\as. set as \ carrier G \ wfactors G as a" by (intro wfactors_exist) simp then obtain as where ascarr[simp]: "set as \ carrier G" and afs: "wfactors G as a" by (auto simp del: carr) have "\as'. set as' \ carrier G \ wfactors G as' a \ length as = length as'" by (metis afs ascarr assms ee_length wfactors_unique) then show "\c. \as'. set as' \ carrier G \ wfactors G as' a \ c = length as'" .. qed lemma (in factorial_monoid) factorcount_unique: assumes afs: "wfactors G as a" and acarr[simp]: "a \ carrier G" and ascarr: "set as \ carrier G" shows "factorcount G a = length as" proof - have "\ac. \as. set as \ carrier G \ wfactors G as a \ ac = length as" by (rule factorcount_exists) simp then obtain ac where alen: "\as. set as \ carrier G \ wfactors G as a \ ac = length as" by auto then have ac: "ac = factorcount G a" unfolding factorcount_def using ascarr by (blast intro: theI2 afs) from ascarr afs have "ac = length as" by (simp add: alen) with ac show ?thesis by simp qed lemma (in factorial_monoid) divides_fcount: assumes dvd: "a divides b" and acarr: "a \ carrier G" and bcarr:"b \ carrier G" shows "factorcount G a \ factorcount G b" proof (rule dividesE[OF dvd]) fix c from assms have "\as. set as \ carrier G \ wfactors G as a" by blast then obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast with acarr have fca: "factorcount G a = length as" by (intro factorcount_unique) assume ccarr: "c \ carrier G" then have "\cs. set cs \ carrier G \ wfactors G cs c" by blast then obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" by blast note [simp] = acarr bcarr ccarr ascarr cscarr assume b: "b = a \ c" from afs cfs have "wfactors G (as@cs) (a \ c)" by (intro wfactors_mult) simp_all with b have "wfactors G (as@cs) b" by simp then have "factorcount G b = length (as@cs)" by (intro factorcount_unique) simp_all then have "factorcount G b = length as + length cs" by simp with fca show ?thesis by simp qed lemma (in factorial_monoid) associated_fcount: assumes acarr: "a \ carrier G" and bcarr: "b \ carrier G" and asc: "a \ b" shows "factorcount G a = factorcount G b" using assms by (auto simp: associated_def factorial_monoid.divides_fcount factorial_monoid_axioms le_antisym) lemma (in factorial_monoid) properfactor_fcount: assumes acarr: "a \ carrier G" and bcarr:"b \ carrier G" and pf: "properfactor G a b" shows "factorcount G a < factorcount G b" proof (rule properfactorE[OF pf], elim dividesE) fix c from assms have "\as. set as \ carrier G \ wfactors G as a" by blast then obtain as where ascarr: "set as \ carrier G" and afs: "wfactors G as a" by blast with acarr have fca: "factorcount G a = length as" by (intro factorcount_unique) assume ccarr: "c \ carrier G" then have "\cs. set cs \ carrier G \ wfactors G cs c" by blast then obtain cs where cscarr: "set cs \ carrier G" and cfs: "wfactors G cs c" by blast assume b: "b = a \ c" have "wfactors G (as@cs) (a \ c)" by (rule wfactors_mult) fact+ with b have "wfactors G (as@cs) b" by simp with ascarr cscarr bcarr have "factorcount G b = length (as@cs)" by (simp add: factorcount_unique) then have fcb: "factorcount G b = length as + length cs" by simp assume nbdvda: "\ b divides a" have "c \ Units G" proof assume cunit:"c \ Units G" have "b \ inv c = a \ c \ inv c" by (simp add: b) also from ccarr acarr cunit have "\ = a \ (c \ inv c)" by (fast intro: m_assoc) also from ccarr cunit have "\ = a \ \" by simp also from acarr have "\ = a" by simp finally have "a = b \ inv c" by simp with ccarr cunit have "b divides a" by (fast intro: dividesI[of "inv c"]) with nbdvda show False by simp qed with cfs have "length cs > 0" by (metis Units_one_closed assoc_unit_r ccarr foldr.simps(1) id_apply length_greater_0_conv wfactors_def) with fca fcb show ?thesis by simp qed sublocale factorial_monoid \ divisor_chain_condition_monoid apply unfold_locales apply (rule wfUNIVI) apply (rule measure_induct[of "factorcount G"]) using properfactor_fcount by auto sublocale factorial_monoid \ primeness_condition_monoid by standard (rule irreducible_prime) lemma (in factorial_monoid) primeness_condition: "primeness_condition_monoid G" .. lemma (in factorial_monoid) gcd_condition [simp]: "gcd_condition_monoid G" by standard (rule gcdof_exists) sublocale factorial_monoid \ gcd_condition_monoid by standard (rule gcdof_exists) lemma (in factorial_monoid) division_weak_lattice [simp]: "weak_lattice (division_rel G)" proof - interpret weak_lower_semilattice "division_rel G" by simp show "weak_lattice (division_rel G)" proof (unfold_locales, simp_all) fix x y assume carr: "x \ carrier G" "y \ carrier G" from lcmof_exists [OF this] obtain z where zcarr: "z \ carrier G" and isgcd: "z lcmof x y" by blast with carr have "least (division_rel G) z (Upper (division_rel G) {x, y})" by (simp add: lcmof_leastUpper[symmetric]) then show "\z. least (division_rel G) z (Upper (division_rel G) {x, y})" by blast qed qed subsection \Factoriality Theorems\ theorem factorial_condition_one: (* Jacobson theorem 2.21 *) "divisor_chain_condition_monoid G \ primeness_condition_monoid G \ factorial_monoid G" proof (rule iffI, clarify) assume dcc: "divisor_chain_condition_monoid G" and pc: "primeness_condition_monoid G" interpret divisor_chain_condition_monoid "G" by (rule dcc) interpret primeness_condition_monoid "G" by (rule pc) show "factorial_monoid G" by (fast intro: factorial_monoidI wfactors_exist wfactors_unique) next assume "factorial_monoid G" then interpret factorial_monoid "G" . show "divisor_chain_condition_monoid G \ primeness_condition_monoid G" by rule unfold_locales qed theorem factorial_condition_two: (* Jacobson theorem 2.22 *) "divisor_chain_condition_monoid G \ gcd_condition_monoid G \ factorial_monoid G" proof (rule iffI, clarify) assume dcc: "divisor_chain_condition_monoid G" and gc: "gcd_condition_monoid G" interpret divisor_chain_condition_monoid "G" by (rule dcc) interpret gcd_condition_monoid "G" by (rule gc) show "factorial_monoid G" by (simp add: factorial_condition_one[symmetric], rule, unfold_locales) next assume "factorial_monoid G" then interpret factorial_monoid "G" . show "divisor_chain_condition_monoid G \ gcd_condition_monoid G" by rule unfold_locales qed end diff --git a/src/HOL/Combinatorics/List_Permutation.thy b/src/HOL/Combinatorics/List_Permutation.thy --- a/src/HOL/Combinatorics/List_Permutation.thy +++ b/src/HOL/Combinatorics/List_Permutation.thy @@ -1,184 +1,142 @@ (* Author: Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker *) section \Permuted Lists\ theory List_Permutation imports Permutations begin text \ Note that multisets already provide the notion of permutated list and hence this theory mostly echoes material already logically present in theory \<^text>\Permutations\; it should be seldom needed. \ -subsection \An inductive definition \ldots\ - -inductive perm :: "'a list \ 'a list \ bool" (infixr \<~~>\ 50) -where - Nil [intro!]: "[] <~~> []" -| swap [intro!]: "y # x # l <~~> x # y # l" -| Cons [intro!]: "xs <~~> ys \ z # xs <~~> z # ys" -| trans [intro]: "xs <~~> ys \ ys <~~> zs \ xs <~~> zs" - -proposition perm_refl [iff]: "l <~~> l" - by (induct l) auto - -text \\ldots that is equivalent to an already existing notion:\ +subsection \An existing notion\ -lemma perm_iff_eq_mset: - \xs <~~> ys \ mset xs = mset ys\ -proof - assume \mset xs = mset ys\ - then show \xs <~~> ys\ - proof (induction xs arbitrary: ys) - case Nil - then show ?case - by simp - next - case (Cons x xs) - from Cons.prems [symmetric] have \mset xs = mset (remove1 x ys)\ - by simp - then have \xs <~~> remove1 x ys\ - by (rule Cons.IH) - then have \x # xs <~~> x # remove1 x ys\ - by (rule perm.Cons) - moreover from Cons.prems have \x \ set ys\ - by (auto dest: union_single_eq_member) - then have \x # remove1 x ys <~~> ys\ - by (induction ys) auto - ultimately show \x # xs <~~> ys\ - by (rule perm.trans) - qed -next - assume \xs <~~> ys\ - then show \mset xs = mset ys\ - by induction simp_all -qed - -theorem mset_eq_perm: \mset xs = mset ys \ xs <~~> ys\ - by (simp add: perm_iff_eq_mset) +abbreviation (input) perm :: \'a list \ 'a list \ bool\ (infixr \<~~>\ 50) + where \xs <~~> ys \ mset xs = mset ys\ subsection \Nontrivial conclusions\ proposition perm_swap: \xs[i := xs ! j, j := xs ! i] <~~> xs\ if \i < length xs\ \j < length xs\ - using that by (simp add: perm_iff_eq_mset mset_swap) + using that by (simp add: mset_swap) proposition mset_le_perm_append: "mset xs \# mset ys \ (\zs. xs @ zs <~~> ys)" - by (auto simp add: perm_iff_eq_mset mset_subset_eq_exists_conv ex_mset dest: sym) + by (auto simp add: mset_subset_eq_exists_conv ex_mset dest: sym) proposition perm_set_eq: "xs <~~> ys \ set xs = set ys" - by (rule mset_eq_setD) (simp add: perm_iff_eq_mset) + by (rule mset_eq_setD) simp proposition perm_distinct_iff: "xs <~~> ys \ distinct xs \ distinct ys" - by (rule mset_eq_imp_distinct_iff) (simp add: perm_iff_eq_mset) + by (rule mset_eq_imp_distinct_iff) simp theorem eq_set_perm_remdups: "set xs = set ys \ remdups xs <~~> remdups ys" - by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq) + by (simp add: set_eq_iff_mset_remdups_eq) proposition perm_remdups_iff_eq_set: "remdups x <~~> remdups y \ set x = set y" - by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq) + by (simp add: set_eq_iff_mset_remdups_eq) theorem permutation_Ex_bij: assumes "xs <~~> ys" shows "\f. bij_betw f {.. (\imset xs = mset ys\ \length xs = length ys\ - by (auto simp add: perm_iff_eq_mset dest: mset_eq_length) + by (auto simp add: dest: mset_eq_length) from \mset xs = mset ys\ obtain p where \p permutes {.. \permute_list p ys = xs\ by (rule mset_eq_permutation) then have \bij_betw p {.. by (simp add: \length xs = length ys\ permutes_imp_bij) moreover have \\i using \permute_list p ys = xs\ \length xs = length ys\ \p permutes {.. permute_list_nth by auto ultimately show ?thesis by blast qed proposition perm_finite: "finite {B. B <~~> A}" - using mset_eq_finite by (auto simp add: perm_iff_eq_mset) + using mset_eq_finite by auto subsection \Trivial conclusions:\ proposition perm_empty_imp: "[] <~~> ys \ ys = []" - by (simp add: perm_iff_eq_mset) + by simp text \\medskip This more general theorem is easier to understand!\ proposition perm_length: "xs <~~> ys \ length xs = length ys" - by (rule mset_eq_length) (simp add: perm_iff_eq_mset) + by (rule mset_eq_length) simp proposition perm_sym: "xs <~~> ys \ ys <~~> xs" - by (simp add: perm_iff_eq_mset) + by simp text \We can insert the head anywhere in the list.\ proposition perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_append_swap: "xs @ ys <~~> ys @ xs" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_append_single: "a # xs <~~> xs @ [a]" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_rev: "rev xs <~~> xs" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_append1: "xs <~~> ys \ l @ xs <~~> l @ ys" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_append2: "xs <~~> ys \ xs @ l <~~> ys @ l" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_empty [iff]: "[] <~~> xs \ xs = []" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_empty2 [iff]: "xs <~~> [] \ xs = []" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_sing_imp: "ys <~~> xs \ xs = [y] \ ys = [y]" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_sing_eq [iff]: "ys <~~> [y] \ ys = [y]" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_sing_eq2 [iff]: "[y] <~~> ys \ ys = [y]" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_remove: "x \ set ys \ ys <~~> x # remove1 x ys" - by (simp add: perm_iff_eq_mset) + by simp text \\medskip Congruence rule\ proposition perm_remove_perm: "xs <~~> ys \ remove1 z xs <~~> remove1 z ys" - by (simp add: perm_iff_eq_mset) + by simp proposition remove_hd [simp]: "remove1 z (z # xs) = xs" - by (simp add: perm_iff_eq_mset) + by simp proposition cons_perm_imp_perm: "z # xs <~~> z # ys \ xs <~~> ys" - by (simp add: perm_iff_eq_mset) + by simp proposition cons_perm_eq [simp]: "z#xs <~~> z#ys \ xs <~~> ys" - by (simp add: perm_iff_eq_mset) + by simp proposition append_perm_imp_perm: "zs @ xs <~~> zs @ ys \ xs <~~> ys" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \ xs <~~> ys" - by (simp add: perm_iff_eq_mset) + by simp proposition perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \ xs <~~> ys" - by (simp add: perm_iff_eq_mset) + by simp end diff --git a/src/HOL/Combinatorics/Permutations.thy b/src/HOL/Combinatorics/Permutations.thy --- a/src/HOL/Combinatorics/Permutations.thy +++ b/src/HOL/Combinatorics/Permutations.thy @@ -1,1764 +1,1769 @@ (* Author: Amine Chaieb, University of Cambridge *) section \Permutations, both general and specifically on finite sets.\ theory Permutations imports "HOL-Library.Multiset" "HOL-Library.Disjoint_Sets" Transposition begin subsection \Auxiliary\ abbreviation (input) fixpoints :: \('a \ 'a) \ 'a set\ where \fixpoints f \ {x. f x = x}\ lemma inj_on_fixpoints: \inj_on f (fixpoints f)\ by (rule inj_onI) simp lemma bij_betw_fixpoints: \bij_betw f (fixpoints f) (fixpoints f)\ using inj_on_fixpoints by (auto simp add: bij_betw_def) subsection \Basic definition and consequences\ definition permutes :: \('a \ 'a) \ 'a set \ bool\ (infixr \permutes\ 41) where \p permutes S \ (\x. x \ S \ p x = x) \ (\y. \!x. p x = y)\ lemma bij_imp_permutes: \p permutes S\ if \bij_betw p S S\ and stable: \\x. x \ S \ p x = x\ proof - note \bij_betw p S S\ moreover have \bij_betw p (- S) (- S)\ by (auto simp add: stable intro!: bij_betw_imageI inj_onI) ultimately have \bij_betw p (S \ - S) (S \ - S)\ by (rule bij_betw_combine) simp then have \\!x. p x = y\ for y by (simp add: bij_iff) with stable show ?thesis by (simp add: permutes_def) qed context fixes p :: \'a \ 'a\ and S :: \'a set\ assumes perm: \p permutes S\ begin lemma permutes_inj: \inj p\ using perm by (auto simp: permutes_def inj_on_def) lemma permutes_image: \p ` S = S\ proof (rule set_eqI) fix x show \x \ p ` S \ x \ S\ proof assume \x \ p ` S\ then obtain y where \y \ S\ \p y = x\ by blast with perm show \x \ S\ by (cases \y = x\) (auto simp add: permutes_def) next assume \x \ S\ with perm obtain y where \y \ S\ \p y = x\ by (metis permutes_def) then show \x \ p ` S\ by blast qed qed lemma permutes_not_in: \x \ S \ p x = x\ using perm by (auto simp: permutes_def) lemma permutes_image_complement: \p ` (- S) = - S\ by (auto simp add: permutes_not_in) lemma permutes_in_image: \p x \ S \ x \ S\ using permutes_image permutes_inj by (auto dest: inj_image_mem_iff) lemma permutes_surj: \surj p\ proof - have \p ` (S \ - S) = p ` S \ p ` (- S)\ by (rule image_Un) then show ?thesis by (simp add: permutes_image permutes_image_complement) qed lemma permutes_inv_o: shows "p \ inv p = id" and "inv p \ p = id" using permutes_inj permutes_surj unfolding inj_iff [symmetric] surj_iff [symmetric] by auto lemma permutes_inverses: shows "p (inv p x) = x" and "inv p (p x) = x" using permutes_inv_o [unfolded fun_eq_iff o_def] by auto lemma permutes_inv_eq: \inv p y = x \ p x = y\ by (auto simp add: permutes_inverses) lemma permutes_inj_on: \inj_on p A\ by (rule inj_on_subset [of _ UNIV]) (auto intro: permutes_inj) lemma permutes_bij: \bij p\ unfolding bij_def by (metis permutes_inj permutes_surj) lemma permutes_imp_bij: \bij_betw p S S\ by (simp add: bij_betw_def permutes_image permutes_inj_on) lemma permutes_subset: \p permutes T\ if \S \ T\ proof (rule bij_imp_permutes) define R where \R = T - S\ with that have \T = R \ S\ \R \ S = {}\ by auto then have \p x = x\ if \x \ R\ for x using that by (auto intro: permutes_not_in) then have \p ` R = R\ by simp with \T = R \ S\ show \bij_betw p T T\ by (simp add: bij_betw_def permutes_inj_on image_Un permutes_image) fix x assume \x \ T\ with \T = R \ S\ show \p x = x\ by (simp add: permutes_not_in) qed lemma permutes_imp_permutes_insert: \p permutes insert x S\ by (rule permutes_subset) auto end lemma permutes_id [simp]: \id permutes S\ by (auto intro: bij_imp_permutes) lemma permutes_empty [simp]: \p permutes {} \ p = id\ proof assume \p permutes {}\ then show \p = id\ by (auto simp add: fun_eq_iff permutes_not_in) next assume \p = id\ then show \p permutes {}\ by simp qed lemma permutes_sing [simp]: \p permutes {a} \ p = id\ proof assume perm: \p permutes {a}\ show \p = id\ proof fix x from perm have \p ` {a} = {a}\ by (rule permutes_image) with perm show \p x = id x\ by (cases \x = a\) (auto simp add: permutes_not_in) qed next assume \p = id\ then show \p permutes {a}\ by simp qed lemma permutes_univ: "p permutes UNIV \ (\y. \!x. p x = y)" by (simp add: permutes_def) lemma permutes_swap_id: "a \ S \ b \ S \ transpose a b permutes S" by (rule bij_imp_permutes) (auto intro: transpose_apply_other) lemma permutes_superset: \p permutes T\ if \p permutes S\ \\x. x \ S - T \ p x = x\ proof - define R U where \R = T \ S\ and \U = S - T\ then have \T = R \ (T - S)\ \S = R \ U\ \R \ U = {}\ by auto from that \U = S - T\ have \p ` U = U\ by simp from \p permutes S\ have \bij_betw p (R \ U) (R \ U)\ by (simp add: permutes_imp_bij \S = R \ U\) moreover have \bij_betw p U U\ using that \U = S - T\ by (simp add: bij_betw_def permutes_inj_on) ultimately have \bij_betw p R R\ using \R \ U = {}\ \R \ U = {}\ by (rule bij_betw_partition) then have \p permutes R\ proof (rule bij_imp_permutes) fix x assume \x \ R\ with \R = T \ S\ \p permutes S\ show \p x = x\ by (cases \x \ S\) (auto simp add: permutes_not_in that(2)) qed then have \p permutes R \ (T - S)\ by (rule permutes_subset) simp with \T = R \ (T - S)\ show ?thesis by simp qed lemma permutes_bij_inv_into: \<^marker>\contributor \Lukas Bulwahn\\ fixes A :: "'a set" and B :: "'b set" assumes "p permutes A" and "bij_betw f A B" shows "(\x. if x \ B then f (p (inv_into A f x)) else x) permutes B" proof (rule bij_imp_permutes) from assms have "bij_betw p A A" "bij_betw f A B" "bij_betw (inv_into A f) B A" by (auto simp add: permutes_imp_bij bij_betw_inv_into) then have "bij_betw (f \ p \ inv_into A f) B B" by (simp add: bij_betw_trans) then show "bij_betw (\x. if x \ B then f (p (inv_into A f x)) else x) B B" by (subst bij_betw_cong[where g="f \ p \ inv_into A f"]) auto next fix x assume "x \ B" then show "(if x \ B then f (p (inv_into A f x)) else x) = x" by auto qed lemma permutes_image_mset: \<^marker>\contributor \Lukas Bulwahn\\ assumes "p permutes A" shows "image_mset p (mset_set A) = mset_set A" using assms by (metis image_mset_mset_set bij_betw_imp_inj_on permutes_imp_bij permutes_image) lemma permutes_implies_image_mset_eq: \<^marker>\contributor \Lukas Bulwahn\\ assumes "p permutes A" "\x. x \ A \ f x = f' (p x)" shows "image_mset f' (mset_set A) = image_mset f (mset_set A)" proof - have "f x = f' (p x)" if "x \# mset_set A" for x using assms(2)[of x] that by (cases "finite A") auto with assms have "image_mset f (mset_set A) = image_mset (f' \ p) (mset_set A)" by (auto intro!: image_mset_cong) also have "\ = image_mset f' (image_mset p (mset_set A))" by (simp add: image_mset.compositionality) also have "\ = image_mset f' (mset_set A)" proof - from assms permutes_image_mset have "image_mset p (mset_set A) = mset_set A" by blast then show ?thesis by simp qed finally show ?thesis .. qed subsection \Group properties\ lemma permutes_compose: "p permutes S \ q permutes S \ q \ p permutes S" unfolding permutes_def o_def by metis lemma permutes_inv: assumes "p permutes S" shows "inv p permutes S" using assms unfolding permutes_def permutes_inv_eq[OF assms] by metis lemma permutes_inv_inv: assumes "p permutes S" shows "inv (inv p) = p" unfolding fun_eq_iff permutes_inv_eq[OF assms] permutes_inv_eq[OF permutes_inv[OF assms]] by blast lemma permutes_invI: assumes perm: "p permutes S" and inv: "\x. x \ S \ p' (p x) = x" and outside: "\x. x \ S \ p' x = x" shows "inv p = p'" proof show "inv p x = p' x" for x proof (cases "x \ S") case True from assms have "p' x = p' (p (inv p x))" by (simp add: permutes_inverses) also from permutes_inv[OF perm] True have "\ = inv p x" by (subst inv) (simp_all add: permutes_in_image) finally show ?thesis .. next case False with permutes_inv[OF perm] show ?thesis by (simp_all add: outside permutes_not_in) qed qed lemma permutes_vimage: "f permutes A \ f -` A = A" by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv]) subsection \Mapping permutations with bijections\ lemma bij_betw_permutations: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A} {\. \ permutes B}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 show ?case using permutes_bij_inv_into[OF _ assms] by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) { fix \ assume "\ permutes B" from permutes_bij_inv_into[OF this bij_inv] and assms have "(\x. if x \ A then inv_into A f (\ (f x)) else x) permutes A" by (simp add: inv_into_inv_into_eq cong: if_cong) } from this show ?case by (auto simp: permutes_inv) next case 1 thus ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed lemma bij_betw_derangements: assumes "bij_betw f A B" shows "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x) {\. \ permutes A \ (\x\A. \ x \ x)} {\. \ permutes B \ (\x\B. \ x \ x)}" (is "bij_betw ?f _ _") proof - let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)" show ?thesis proof (rule bij_betw_byWitness [of _ ?g], goal_cases) case 3 have "?f \ x \ x" if "\ permutes A" "\x. x \ A \ \ x \ x" "x \ B" for \ x using that and assms by (metis bij_betwE bij_betw_imp_inj_on bij_betw_imp_surj_on inv_into_f_f inv_into_into permutes_imp_bij) with permutes_bij_inv_into[OF _ assms] show ?case by auto next case 4 have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) have "?g \ permutes A" if "\ permutes B" for \ using permutes_bij_inv_into[OF that bij_inv] and assms by (simp add: inv_into_inv_into_eq cong: if_cong) moreover have "?g \ x \ x" if "\ permutes B" "\x. x \ B \ \ x \ x" "x \ A" for \ x using that and assms by (metis bij_betwE bij_betw_imp_surj_on f_inv_into_f permutes_imp_bij) ultimately show ?case by auto next case 1 thus ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left dest: bij_betwE) next case 2 moreover have "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms) ultimately show ?case using assms by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right dest: bij_betwE) qed qed subsection \The number of permutations on a finite set\ lemma permutes_insert_lemma: assumes "p permutes (insert a S)" shows "transpose a (p a) \ p permutes S" apply (rule permutes_superset[where S = "insert a S"]) apply (rule permutes_compose[OF assms]) apply (rule permutes_swap_id, simp) using permutes_in_image[OF assms, of a] apply simp apply (auto simp add: Ball_def) done lemma permutes_insert: "{p. p permutes (insert a S)} = (\(b, p). transpose a b \ p) ` {(b, p). b \ insert a S \ p \ {p. p permutes S}}" proof - have "p permutes insert a S \ (\b q. p = transpose a b \ q \ b \ insert a S \ q permutes S)" for p proof - have "\b q. p = transpose a b \ q \ b \ insert a S \ q permutes S" if p: "p permutes insert a S" proof - let ?b = "p a" let ?q = "transpose a (p a) \ p" have *: "p = transpose a ?b \ ?q" by (simp add: fun_eq_iff o_assoc) have **: "?b \ insert a S" unfolding permutes_in_image[OF p] by simp from permutes_insert_lemma[OF p] * ** show ?thesis by blast qed moreover have "p permutes insert a S" if bq: "p = transpose a b \ q" "b \ insert a S" "q permutes S" for b q proof - from permutes_subset[OF bq(3), of "insert a S"] have q: "q permutes insert a S" by auto have a: "a \ insert a S" by simp from bq(1) permutes_compose[OF q permutes_swap_id[OF a bq(2)]] show ?thesis by simp qed ultimately show ?thesis by blast qed then show ?thesis by auto qed lemma card_permutations: assumes "card S = n" and "finite S" shows "card {p. p permutes S} = fact n" using assms(2,1) proof (induct arbitrary: n) case empty then show ?case by simp next case (insert x F) { fix n assume card_insert: "card (insert x F) = n" let ?xF = "{p. p permutes insert x F}" let ?pF = "{p. p permutes F}" let ?pF' = "{(b, p). b \ insert x F \ p \ ?pF}" let ?g = "(\(b, p). transpose x b \ p)" have xfgpF': "?xF = ?g ` ?pF'" by (rule permutes_insert[of x F]) from \x \ F\ \finite F\ card_insert have Fs: "card F = n - 1" by auto from \finite F\ insert.hyps Fs have pFs: "card ?pF = fact (n - 1)" by auto then have "finite ?pF" by (auto intro: card_ge_0_finite) with \finite F\ card.insert_remove have pF'f: "finite ?pF'" apply (simp only: Collect_case_prod Collect_mem_eq) apply (rule finite_cartesian_product) apply simp_all done have ginj: "inj_on ?g ?pF'" proof - { fix b p c q assume bp: "(b, p) \ ?pF'" assume cq: "(c, q) \ ?pF'" assume eq: "?g (b, p) = ?g (c, q)" from bp cq have pF: "p permutes F" and qF: "q permutes F" by auto from pF \x \ F\ eq have "b = ?g (b, p) x" by (auto simp: permutes_def fun_upd_def fun_eq_iff) also from qF \x \ F\ eq have "\ = ?g (c, q) x" by (auto simp: fun_upd_def fun_eq_iff) also from qF \x \ F\ have "\ = c" by (auto simp: permutes_def fun_upd_def fun_eq_iff) finally have "b = c" . then have "transpose x b = transpose x c" by simp with eq have "transpose x b \ p = transpose x b \ q" by simp then have "transpose x b \ (transpose x b \ p) = transpose x b \ (transpose x b \ q)" by simp then have "p = q" by (simp add: o_assoc) with \b = c\ have "(b, p) = (c, q)" by simp } then show ?thesis unfolding inj_on_def by blast qed from \x \ F\ \finite F\ card_insert have "n \ 0" by auto then have "\m. n = Suc m" by presburger then obtain m where n: "n = Suc m" by blast from pFs card_insert have *: "card ?xF = fact n" unfolding xfgpF' card_image[OF ginj] using \finite F\ \finite ?pF\ by (simp only: Collect_case_prod Collect_mem_eq card_cartesian_product) (simp add: n) from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" by (simp add: xfgpF' n) from * have "card ?xF = fact n" unfolding xFf by blast } with insert show ?case by simp qed lemma finite_permutations: assumes "finite S" shows "finite {p. p permutes S}" using card_permutations[OF refl assms] by (auto intro: card_ge_0_finite) subsection \Hence a sort of induction principle composing by swaps\ lemma permutes_induct [consumes 2, case_names id swap]: \P p\ if \p permutes S\ \finite S\ and id: \P id\ and swap: \\a b p. a \ S \ b \ S \ p permutes S \ P p \ P (transpose a b \ p)\ using \finite S\ \p permutes S\ swap proof (induction S arbitrary: p) case empty with id show ?case by (simp only: permutes_empty) next case (insert x S p) define q where \q = transpose x (p x) \ p\ then have swap_q: \transpose x (p x) \ q = p\ by (simp add: o_assoc) from \p permutes insert x S\ have \q permutes S\ by (simp add: q_def permutes_insert_lemma) then have \q permutes insert x S\ by (simp add: permutes_imp_permutes_insert) from \q permutes S\ have \P q\ by (auto intro: insert.IH insert.prems(2) permutes_imp_permutes_insert) have \x \ insert x S\ by simp moreover from \p permutes insert x S\ have \p x \ insert x S\ using permutes_in_image [of p \insert x S\ x] by simp ultimately have \P (transpose x (p x) \ q)\ using \q permutes insert x S\ \P q\ by (rule insert.prems(2)) then show ?case by (simp add: swap_q) qed lemma permutes_rev_induct [consumes 2, case_names id swap]: \P p\ if \p permutes S\ \finite S\ and id': \P id\ and swap': \\a b p. a \ S \ b \ S \ p permutes S \ P p \ P (p \ transpose a b)\ using \p permutes S\ \finite S\ proof (induction rule: permutes_induct) case id from id' show ?case . next case (swap a b p) then have \bij p\ using permutes_bij by blast have \P (p \ transpose (inv p a) (inv p b))\ by (rule swap') (auto simp add: swap permutes_in_image permutes_inv) also have \p \ transpose (inv p a) (inv p b) = transpose a b \ p\ using \bij p\ by (rule transpose_comp_eq [symmetric]) finally show ?case . qed subsection \Permutations of index set for iterated operations\ lemma (in comm_monoid_set) permute: assumes "p permutes S" shows "F g S = F (g \ p) S" proof - from \p permutes S\ have "inj p" by (rule permutes_inj) then have "inj_on p S" by (auto intro: subset_inj_on) then have "F g (p ` S) = F (g \ p) S" by (rule reindex) moreover from \p permutes S\ have "p ` S = S" by (rule permutes_image) ultimately show ?thesis by simp qed subsection \Permutations as transposition sequences\ inductive swapidseq :: "nat \ ('a \ 'a) \ bool" where id[simp]: "swapidseq 0 id" | comp_Suc: "swapidseq n p \ a \ b \ swapidseq (Suc n) (transpose a b \ p)" declare id[unfolded id_def, simp] definition "permutation p \ (\n. swapidseq n p)" subsection \Some closure properties of the set of permutations, with lengths\ lemma permutation_id[simp]: "permutation id" unfolding permutation_def by (rule exI[where x=0]) simp declare permutation_id[unfolded id_def, simp] lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (transpose a b)" apply clarsimp using comp_Suc[of 0 id a b] apply simp done lemma permutation_swap_id: "permutation (transpose a b)" proof (cases "a = b") case True then show ?thesis by simp next case False then show ?thesis unfolding permutation_def using swapidseq_swap[of a b] by blast qed lemma swapidseq_comp_add: "swapidseq n p \ swapidseq m q \ swapidseq (n + m) (p \ q)" proof (induct n p arbitrary: m q rule: swapidseq.induct) case (id m q) then show ?case by simp next case (comp_Suc n p a b m q) have eq: "Suc n + m = Suc (n + m)" by arith show ?case apply (simp only: eq comp_assoc) apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3) apply blast+ done qed lemma permutation_compose: "permutation p \ permutation q \ permutation (p \ q)" unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis lemma swapidseq_endswap: "swapidseq n p \ a \ b \ swapidseq (Suc n) (p \ transpose a b)" by (induct n p rule: swapidseq.induct) (use swapidseq_swap[of a b] in \auto simp add: comp_assoc intro: swapidseq.comp_Suc\) lemma swapidseq_inverse_exists: "swapidseq n p \ \q. swapidseq n q \ p \ q = id \ q \ p = id" proof (induct n p rule: swapidseq.induct) case id then show ?case by (rule exI[where x=id]) simp next case (comp_Suc n p a b) from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast let ?q = "q \ transpose a b" note H = comp_Suc.hyps from swapidseq_swap[of a b] H(3) have *: "swapidseq 1 (transpose a b)" by simp from swapidseq_comp_add[OF q(1) *] have **: "swapidseq (Suc n) ?q" by simp have "transpose a b \ p \ ?q = transpose a b \ (p \ q) \ transpose a b" by (simp add: o_assoc) also have "\ = id" by (simp add: q(2)) finally have ***: "transpose a b \ p \ ?q = id" . have "?q \ (transpose a b \ p) = q \ (transpose a b \ transpose a b) \ p" by (simp only: o_assoc) then have "?q \ (transpose a b \ p) = id" by (simp add: q(3)) with ** *** show ?case by blast qed lemma swapidseq_inverse: assumes "swapidseq n p" shows "swapidseq n (inv p)" using swapidseq_inverse_exists[OF assms] inv_unique_comp[of p] by auto lemma permutation_inverse: "permutation p \ permutation (inv p)" using permutation_def swapidseq_inverse by blast subsection \Various combinations of transpositions with 2, 1 and 0 common elements\ lemma swap_id_common:" a \ c \ b \ c \ transpose a b \ transpose a c = transpose b c \ transpose a b" by (simp add: fun_eq_iff transpose_def) lemma swap_id_common': "a \ b \ a \ c \ transpose a c \ transpose b c = transpose b c \ transpose a b" by (simp add: fun_eq_iff transpose_def) lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \ transpose a b \ transpose c d = transpose c d \ transpose a b" by (simp add: fun_eq_iff transpose_def) subsection \The identity map only has even transposition sequences\ lemma symmetry_lemma: assumes "\a b c d. P a b c d \ P a b d c" and "\a b c d. a \ b \ c \ d \ a = c \ b = d \ a = c \ b \ d \ a \ c \ b = d \ a \ c \ a \ d \ b \ c \ b \ d \ P a b c d" shows "\a b c d. a \ b \ c \ d \ P a b c d" using assms by metis lemma swap_general: "a \ b \ c \ d \ transpose a b \ transpose c d = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ transpose a b \ transpose c d = transpose x y \ transpose a z)" proof - assume neq: "a \ b" "c \ d" have "a \ b \ c \ d \ (transpose a b \ transpose c d = id \ (\x y z. x \ a \ y \ a \ z \ a \ x \ y \ transpose a b \ transpose c d = transpose x y \ transpose a z))" apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d]) apply (simp_all only: ac_simps) apply (metis id_comp swap_id_common swap_id_common' swap_id_independent transpose_comp_involutory) done with neq show ?thesis by metis qed lemma swapidseq_id_iff[simp]: "swapidseq 0 p \ p = id" using swapidseq.cases[of 0 p "p = id"] by auto lemma swapidseq_cases: "swapidseq n p \ n = 0 \ p = id \ (\a b q m. n = Suc m \ p = transpose a b \ q \ swapidseq m q \ a \ b)" apply (rule iffI) apply (erule swapidseq.cases[of n p]) apply simp apply (rule disjI2) apply (rule_tac x= "a" in exI) apply (rule_tac x= "b" in exI) apply (rule_tac x= "pa" in exI) apply (rule_tac x= "na" in exI) apply simp apply auto apply (rule comp_Suc, simp_all) done lemma fixing_swapidseq_decrease: assumes "swapidseq n p" and "a \ b" and "(transpose a b \ p) a = a" shows "n \ 0 \ swapidseq (n - 1) (transpose a b \ p)" using assms proof (induct n arbitrary: p a b) case 0 then show ?case by (auto simp add: fun_upd_def) next case (Suc n p a b) from Suc.prems(1) swapidseq_cases[of "Suc n" p] obtain c d q m where cdqm: "Suc n = Suc m" "p = transpose c d \ q" "swapidseq m q" "c \ d" "n = m" by auto consider "transpose a b \ transpose c d = id" | x y z where "x \ a" "y \ a" "z \ a" "x \ y" "transpose a b \ transpose c d = transpose x y \ transpose a z" using swap_general[OF Suc.prems(2) cdqm(4)] by metis then show ?case proof cases case 1 then show ?thesis by (simp only: cdqm o_assoc) (simp add: cdqm) next case prems: 2 then have az: "a \ z" by simp from prems have *: "(transpose x y \ h) a = a \ h a = a" for h by (simp add: transpose_def) from cdqm(2) have "transpose a b \ p = transpose a b \ (transpose c d \ q)" by simp then have "transpose a b \ p = transpose x y \ (transpose a z \ q)" by (simp add: o_assoc prems) then have "(transpose a b \ p) a = (transpose x y \ (transpose a z \ q)) a" by simp then have "(transpose x y \ (transpose a z \ q)) a = a" unfolding Suc by metis then have "(transpose a z \ q) a = a" by (simp only: *) from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az this] have **: "swapidseq (n - 1) (transpose a z \ q)" "n \ 0" by blast+ from \n \ 0\ have ***: "Suc n - 1 = Suc (n - 1)" by auto show ?thesis apply (simp only: cdqm(2) prems o_assoc ***) apply (simp only: Suc_not_Zero simp_thms comp_assoc) apply (rule comp_Suc) using ** prems apply blast+ done qed qed lemma swapidseq_identity_even: assumes "swapidseq n (id :: 'a \ 'a)" shows "even n" using \swapidseq n id\ proof (induct n rule: nat_less_induct) case H: (1 n) consider "n = 0" | a b :: 'a and q m where "n = Suc m" "id = transpose a b \ q" "swapidseq m q" "a \ b" using H(2)[unfolded swapidseq_cases[of n id]] by auto then show ?case proof cases case 1 then show ?thesis by presburger next case h: 2 from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]] have m: "m \ 0" "swapidseq (m - 1) (id :: 'a \ 'a)" by auto from h m have mn: "m - 1 < n" by arith from H(1)[rule_format, OF mn m(2)] h(1) m(1) show ?thesis by presburger qed qed subsection \Therefore we have a welldefined notion of parity\ definition "evenperm p = even (SOME n. swapidseq n p)" lemma swapidseq_even_even: assumes m: "swapidseq m p" and n: "swapidseq n p" shows "even m \ even n" proof - from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]] show ?thesis by arith qed lemma evenperm_unique: assumes p: "swapidseq n p" and n:"even n = b" shows "evenperm p = b" unfolding n[symmetric] evenperm_def apply (rule swapidseq_even_even[where p = p]) apply (rule someI[where x = n]) using p apply blast+ done subsection \And it has the expected composition properties\ lemma evenperm_id[simp]: "evenperm id = True" by (rule evenperm_unique[where n = 0]) simp_all lemma evenperm_identity [simp]: \evenperm (\x. x)\ using evenperm_id by (simp add: id_def [abs_def]) lemma evenperm_swap: "evenperm (transpose a b) = (a = b)" by (rule evenperm_unique[where n="if a = b then 0 else 1"]) (simp_all add: swapidseq_swap) lemma evenperm_comp: assumes "permutation p" "permutation q" shows "evenperm (p \ q) \ evenperm p = evenperm q" proof - from assms obtain n m where n: "swapidseq n p" and m: "swapidseq m q" unfolding permutation_def by blast have "even (n + m) \ (even n \ even m)" by arith from evenperm_unique[OF n refl] evenperm_unique[OF m refl] and evenperm_unique[OF swapidseq_comp_add[OF n m] this] show ?thesis by blast qed lemma evenperm_inv: assumes "permutation p" shows "evenperm (inv p) = evenperm p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast show ?thesis by (rule evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]]) qed subsection \A more abstract characterization of permutations\ lemma permutation_bijective: assumes "permutation p" shows "bij p" proof - from assms obtain n where n: "swapidseq n p" unfolding permutation_def by blast from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id" by blast then show ?thesis unfolding bij_iff apply (auto simp add: fun_eq_iff) apply metis done qed lemma permutation_finite_support: assumes "permutation p" shows "finite {x. p x \ x}" proof - from assms obtain n where "swapidseq n p" unfolding permutation_def by blast then show ?thesis proof (induct n p rule: swapidseq.induct) case id then show ?case by simp next case (comp_Suc n p a b) let ?S = "insert a (insert b {x. p x \ x})" from comp_Suc.hyps(2) have *: "finite ?S" by simp from \a \ b\ have **: "{x. (transpose a b \ p) x \ x} \ ?S" by auto show ?case by (rule finite_subset[OF ** *]) qed qed lemma permutation_lemma: assumes "finite S" and "bij p" and "\x. x \ S \ p x = x" shows "permutation p" using assms proof (induct S arbitrary: p rule: finite_induct) case empty then show ?case by simp next case (insert a F p) let ?r = "transpose a (p a) \ p" let ?q = "transpose a (p a) \ ?r" have *: "?r a = a" by simp from insert * have **: "\x. x \ F \ ?r x = x" by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3)) have "bij ?r" using insert by (simp add: bij_comp) have "permutation ?r" by (rule insert(3)[OF \bij ?r\ **]) then have "permutation ?q" by (simp add: permutation_compose permutation_swap_id) then show ?case by (simp add: o_assoc) qed lemma permutation: "permutation p \ bij p \ finite {x. p x \ x}" (is "?lhs \ ?b \ ?f") proof assume ?lhs with permutation_bijective permutation_finite_support show "?b \ ?f" by auto next assume "?b \ ?f" then have "?f" "?b" by blast+ from permutation_lemma[OF this] show ?lhs by blast qed lemma permutation_inverse_works: assumes "permutation p" shows "inv p \ p = id" and "p \ inv p = id" using permutation_bijective [OF assms] by (auto simp: bij_def inj_iff surj_iff) lemma permutation_inverse_compose: assumes p: "permutation p" and q: "permutation q" shows "inv (p \ q) = inv q \ inv p" proof - note ps = permutation_inverse_works[OF p] note qs = permutation_inverse_works[OF q] have "p \ q \ (inv q \ inv p) = p \ (q \ inv q) \ inv p" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have *: "p \ q \ (inv q \ inv p) = id" . have "inv q \ inv p \ (p \ q) = inv q \ (inv p \ p) \ q" by (simp add: o_assoc) also have "\ = id" by (simp add: ps qs) finally have **: "inv q \ inv p \ (p \ q) = id" . show ?thesis by (rule inv_unique_comp[OF * **]) qed subsection \Relation to \permutes\\ lemma permutes_imp_permutation: \permutation p\ if \finite S\ \p permutes S\ proof - from \p permutes S\ have \{x. p x \ x} \ S\ by (auto dest: permutes_not_in) then have \finite {x. p x \ x}\ using \finite S\ by (rule finite_subset) moreover from \p permutes S\ have \bij p\ by (auto dest: permutes_bij) ultimately show ?thesis by (simp add: permutation) qed lemma permutation_permutesE: assumes \permutation p\ obtains S where \finite S\ \p permutes S\ proof - from assms have fin: \finite {x. p x \ x}\ by (simp add: permutation) from assms have \bij p\ by (simp add: permutation) also have \UNIV = {x. p x \ x} \ {x. p x = x}\ by auto finally have \bij_betw p {x. p x \ x} {x. p x \ x}\ by (rule bij_betw_partition) (auto simp add: bij_betw_fixpoints) then have \p permutes {x. p x \ x}\ by (auto intro: bij_imp_permutes) with fin show thesis .. qed lemma permutation_permutes: "permutation p \ (\S. finite S \ p permutes S)" by (auto elim: permutation_permutesE intro: permutes_imp_permutation) subsection \Sign of a permutation as a real number\ definition sign :: \('a \ 'a) \ int\ \ \TODO: prefer less generic name\ where \sign p = (if evenperm p then 1 else - 1)\ lemma sign_cases [case_names even odd]: obtains \sign p = 1\ | \sign p = - 1\ by (cases \evenperm p\) (simp_all add: sign_def) lemma sign_nz [simp]: "sign p \ 0" by (cases p rule: sign_cases) simp_all lemma sign_id [simp]: "sign id = 1" by (simp add: sign_def) lemma sign_identity [simp]: \sign (\x. x) = 1\ by (simp add: sign_def) lemma sign_inverse: "permutation p \ sign (inv p) = sign p" by (simp add: sign_def evenperm_inv) lemma sign_compose: "permutation p \ permutation q \ sign (p \ q) = sign p * sign q" by (simp add: sign_def evenperm_comp) lemma sign_swap_id: "sign (transpose a b) = (if a = b then 1 else - 1)" by (simp add: sign_def evenperm_swap) lemma sign_idempotent [simp]: "sign p * sign p = 1" by (simp add: sign_def) lemma sign_left_idempotent [simp]: \sign p * (sign p * sign q) = sign q\ by (simp add: sign_def) term "(bij, bij_betw, permutation)" subsection \Permuting a list\ text \This function permutes a list by applying a permutation to the indices.\ definition permute_list :: "(nat \ nat) \ 'a list \ 'a list" where "permute_list f xs = map (\i. xs ! (f i)) [0.. g) xs = permute_list g (permute_list f xs)" using assms[THEN permutes_in_image] by (auto simp add: permute_list_def) lemma permute_list_ident [simp]: "permute_list (\x. x) xs = xs" by (simp add: permute_list_def map_nth) lemma permute_list_id [simp]: "permute_list id xs = xs" by (simp add: id_def) lemma mset_permute_list [simp]: fixes xs :: "'a list" assumes "f permutes {.. x < length xs" for x using permutes_in_image[OF assms] by auto have "count (mset (permute_list f xs)) y = card ((\i. xs ! f i) -` {y} \ {..i. xs ! f i) -` {y} \ {.. y = xs ! i}" by auto also from assms have "card \ = card {i. i < length xs \ y = xs ! i}" by (intro card_vimage_inj) (auto simp: permutes_inj permutes_surj) also have "\ = count (mset xs) y" by (simp add: count_mset length_filter_conv_card) finally show "count (mset (permute_list f xs)) y = count (mset xs) y" by simp qed lemma set_permute_list [simp]: assumes "f permutes {.. i < length ys" for i by simp have "permute_list f (zip xs ys) = map (\i. zip xs ys ! f i) [0.. = map (\(x, y). (xs ! f x, ys ! f y)) (zip [0.. = zip (permute_list f xs) (permute_list f ys)" by (simp_all add: permute_list_def zip_map_map) finally show ?thesis . qed lemma map_of_permute: assumes "\ permutes fst ` set xs" shows "map_of xs \ \ = map_of (map (\(x,y). (inv \ x, y)) xs)" (is "_ = map_of (map ?f _)") proof from assms have "inj \" "surj \" by (simp_all add: permutes_inj permutes_surj) then show "(map_of xs \ \) x = map_of (map ?f xs) x" for x by (induct xs) (auto simp: inv_f_f surj_f_inv_f) qed +lemma list_all2_permute_list_iff: + \list_all2 P (permute_list p xs) (permute_list p ys) \ list_all2 P xs ys\ + if \p permutes {.. + using that by (auto simp add: list_all2_iff simp flip: permute_list_zip) + subsection \More lemmas about permutations\ lemma permutes_in_funpow_image: \<^marker>\contributor \Lars Noschinski\\ assumes "f permutes S" "x \ S" shows "(f ^^ n) x \ S" using assms by (induction n) (auto simp: permutes_in_image) lemma permutation_self: \<^marker>\contributor \Lars Noschinski\\ assumes \permutation p\ obtains n where \n > 0\ \(p ^^ n) x = x\ proof (cases \p x = x\) case True with that [of 1] show thesis by simp next case False from \permutation p\ have \inj p\ by (intro permutation_bijective bij_is_inj) moreover from \p x \ x\ have \(p ^^ Suc n) x \ (p ^^ n) x\ for n proof (induction n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have "p (p x) \ p x" proof (rule notI) assume "p (p x) = p x" then show False using \p x \ x\ \inj p\ by (simp add: inj_eq) qed have "(p ^^ Suc (Suc n)) x = (p ^^ Suc n) (p x)" by (simp add: funpow_swap1) also have "\ \ (p ^^ n) (p x)" by (rule Suc) fact also have "(p ^^ n) (p x) = (p ^^ Suc n) x" by (simp add: funpow_swap1) finally show ?case by simp qed then have "{y. \n. y = (p ^^ n) x} \ {x. p x \ x}" by auto then have "finite {y. \n. y = (p ^^ n) x}" using permutation_finite_support[OF assms] by (rule finite_subset) ultimately obtain n where \n > 0\ \(p ^^ n) x = x\ by (rule funpow_inj_finite) with that [of n] show thesis by blast qed text \The following few lemmas were contributed by Lukas Bulwahn.\ lemma count_image_mset_eq_card_vimage: assumes "finite A" shows "count (image_mset f (mset_set A)) b = card {a \ A. f a = b}" using assms proof (induct A) case empty show ?case by simp next case (insert x F) show ?case proof (cases "f x = b") case True with insert.hyps have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \ F. f a = f x})" by auto also from insert.hyps(1,2) have "\ = card (insert x {a \ F. f a = f x})" by simp also from \f x = b\ have "card (insert x {a \ F. f a = f x}) = card {a \ insert x F. f a = b}" by (auto intro: arg_cong[where f="card"]) finally show ?thesis using insert by auto next case False then have "{a \ F. f a = b} = {a \ insert x F. f a = b}" by auto with insert False show ?thesis by simp qed qed \ \Prove \image_mset_eq_implies_permutes\ ...\ lemma image_mset_eq_implies_permutes: fixes f :: "'a \ 'b" assumes "finite A" and mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A)" obtains p where "p permutes A" and "\x\A. f x = f' (p x)" proof - from \finite A\ have [simp]: "finite {a \ A. f a = (b::'b)}" for f b by auto have "f ` A = f' ` A" proof - from \finite A\ have "f ` A = f ` (set_mset (mset_set A))" by simp also have "\ = f' ` set_mset (mset_set A)" by (metis mset_eq multiset.set_map) also from \finite A\ have "\ = f' ` A" by simp finally show ?thesis . qed have "\b\(f ` A). \p. bij_betw p {a \ A. f a = b} {a \ A. f' a = b}" proof fix b from mset_eq have "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A)) b" by simp with \finite A\ have "card {a \ A. f a = b} = card {a \ A. f' a = b}" by (simp add: count_image_mset_eq_card_vimage) then show "\p. bij_betw p {a\A. f a = b} {a \ A. f' a = b}" by (intro finite_same_card_bij) simp_all qed then have "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" by (rule bchoice) then obtain p where p: "\b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" .. define p' where "p' = (\a. if a \ A then p (f a) a else a)" have "p' permutes A" proof (rule bij_imp_permutes) have "disjoint_family_on (\i. {a \ A. f' a = i}) (f ` A)" by (auto simp: disjoint_family_on_def) moreover have "bij_betw (\a. p (f a) a) {a \ A. f a = b} {a \ A. f' a = b}" if "b \ f ` A" for b using p that by (subst bij_betw_cong[where g="p b"]) auto ultimately have "bij_betw (\a. p (f a) a) (\b\f ` A. {a \ A. f a = b}) (\b\f ` A. {a \ A. f' a = b})" by (rule bij_betw_UNION_disjoint) moreover have "(\b\f ` A. {a \ A. f a = b}) = A" by auto moreover from \f ` A = f' ` A\ have "(\b\f ` A. {a \ A. f' a = b}) = A" by auto ultimately show "bij_betw p' A A" unfolding p'_def by (subst bij_betw_cong[where g="(\a. p (f a) a)"]) auto next show "\x. x \ A \ p' x = x" by (simp add: p'_def) qed moreover from p have "\x\A. f x = f' (p' x)" unfolding p'_def using bij_betwE by fastforce ultimately show ?thesis .. qed \ \... and derive the existing property:\ lemma mset_eq_permutation: fixes xs ys :: "'a list" assumes mset_eq: "mset xs = mset ys" obtains p where "p permutes {..i. xs ! i) (mset_set {..i. ys ! i) (mset_set {..i\{..i \ S. p i \ i" shows "p = id" proof - have "p n = n" for n using assms proof (induct n arbitrary: S rule: less_induct) case (less n) show ?case proof (cases "n \ S") case False with less(2) show ?thesis unfolding permutes_def by metis next case True with less(3) have "p n < n \ p n = n" by auto then show ?thesis proof assume "p n < n" with less have "p (p n) = p n" by metis with permutes_inj[OF less(2)] have "p n = n" unfolding inj_def by blast with \p n < n\ have False by simp then show ?thesis .. qed qed qed then show ?thesis by (auto simp: fun_eq_iff) qed lemma permutes_natset_ge: fixes S :: "'a::wellorder set" assumes p: "p permutes S" and le: "\i \ S. p i \ i" shows "p = id" proof - have "i \ inv p i" if "i \ S" for i proof - from that permutes_in_image[OF permutes_inv[OF p]] have "inv p i \ S" by simp with le have "p (inv p i) \ inv p i" by blast with permutes_inverses[OF p] show ?thesis by simp qed then have "\i\S. inv p i \ i" by blast from permutes_natset_le[OF permutes_inv[OF p] this] have "inv p = inv id" by simp then show ?thesis apply (subst permutes_inv_inv[OF p, symmetric]) apply (rule inv_unique_comp) apply simp_all done qed lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto using permutes_inv_inv permutes_inv apply auto apply (rule_tac x="inv x" in exI) apply auto done lemma image_compose_permutations_left: assumes "q permutes S" shows "{q \ p |p. p permutes S} = {p. p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "inv q \ x" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o) done lemma image_compose_permutations_right: assumes "q permutes S" shows "{p \ q | p. p permutes S} = {p . p permutes S}" apply (rule set_eqI) apply auto apply (rule permutes_compose) using assms apply auto apply (rule_tac x = "x \ inv q" in exI) apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o comp_assoc) done lemma permutes_in_seg: "p permutes {1 ..n} \ i \ {1..n} \ 1 \ p i \ p i \ n" by (simp add: permutes_def) metis lemma sum_permutations_inverse: "sum f {p. p permutes S} = sum (\p. f(inv p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p . p permutes S}" have *: "inj_on inv ?S" proof (auto simp add: inj_on_def) fix q r assume q: "q permutes S" and r: "r permutes S" and qr: "inv q = inv r" then have "inv (inv q) = inv (inv r)" by simp with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r" by metis qed have **: "inv ` ?S = ?S" using image_inverse_permutations by blast have ***: "?rhs = sum (f \ inv) ?S" by (simp add: o_def) from sum.reindex[OF *, of f] show ?thesis by (simp only: ** ***) qed lemma setum_permutations_compose_left: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(q \ p)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ ((\) q)) ?S" by (simp add: o_def) have **: "inj_on ((\) q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "q \ p = q \ r" then have "inv q \ q \ p = inv q \ q \ r" by (simp add: comp_assoc) with permutes_inj[OF q, unfolded inj_iff] show "p = r" by simp qed have "((\) q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed lemma sum_permutations_compose_right: assumes q: "q permutes S" shows "sum f {p. p permutes S} = sum (\p. f(p \ q)) {p. p permutes S}" (is "?lhs = ?rhs") proof - let ?S = "{p. p permutes S}" have *: "?rhs = sum (f \ (\p. p \ q)) ?S" by (simp add: o_def) have **: "inj_on (\p. p \ q) ?S" proof (auto simp add: inj_on_def) fix p r assume "p permutes S" and r: "r permutes S" and rp: "p \ q = r \ q" then have "p \ (q \ inv q) = r \ (q \ inv q)" by (simp add: o_assoc) with permutes_surj[OF q, unfolded surj_iff] show "p = r" by simp qed from image_compose_permutations_right[OF q] have "(\p. p \ q) ` ?S = ?S" by auto with * sum.reindex[OF **, of f] show ?thesis by (simp only:) qed lemma inv_inj_on_permutes: \inj_on inv {p. p permutes S}\ proof (intro inj_onI, unfold mem_Collect_eq) fix p q assume p: "p permutes S" and q: "q permutes S" and eq: "inv p = inv q" have "inv (inv p) = inv (inv q)" using eq by simp thus "p = q" using inv_inv_eq[OF permutes_bij] p q by metis qed lemma permutes_pair_eq: \{(p s, s) |s. s \ S} = {(s, inv p s) |s. s \ S}\ (is \?L = ?R\) if \p permutes S\ proof show "?L \ ?R" proof fix x assume "x \ ?L" then obtain s where x: "x = (p s, s)" and s: "s \ S" by auto note x also have "(p s, s) = (p s, Hilbert_Choice.inv p (p s))" using permutes_inj [OF that] inv_f_f by auto also have "... \ ?R" using s permutes_in_image[OF that] by auto finally show "x \ ?R". qed show "?R \ ?L" proof fix x assume "x \ ?R" then obtain s where x: "x = (s, Hilbert_Choice.inv p s)" (is "_ = (s, ?ips)") and s: "s \ S" by auto note x also have "(s, ?ips) = (p ?ips, ?ips)" using inv_f_f[OF permutes_inj[OF permutes_inv[OF that]]] using inv_inv_eq[OF permutes_bij[OF that]] by auto also have "... \ ?L" using s permutes_in_image[OF permutes_inv[OF that]] by auto finally show "x \ ?L". qed qed context fixes p and n i :: nat assumes p: \p permutes {0.. and i: \i < n\ begin lemma permutes_nat_less: \p i < n\ proof - have \?thesis \ p i \ {0.. by simp also from p have \p i \ {0.. i \ {0.. by (rule permutes_in_image) finally show ?thesis using i by simp qed lemma permutes_nat_inv_less: \inv p i < n\ proof - from p have \inv p permutes {0.. by (rule permutes_inv) then show ?thesis using i by (rule Permutations.permutes_nat_less) qed end context comm_monoid_set begin lemma permutes_inv: \F (\s. g (p s) s) S = F (\s. g s (inv p s)) S\ (is \?l = ?r\) if \p permutes S\ proof - let ?g = "\(x, y). g x y" let ?ps = "\s. (p s, s)" let ?ips = "\s. (s, inv p s)" have inj1: "inj_on ?ps S" by (rule inj_onI) auto have inj2: "inj_on ?ips S" by (rule inj_onI) auto have "?l = F ?g (?ps ` S)" using reindex [OF inj1, of ?g] by simp also have "?ps ` S = {(p s, s) |s. s \ S}" by auto also have "... = {(s, inv p s) |s. s \ S}" unfolding permutes_pair_eq [OF that] by simp also have "... = ?ips ` S" by auto also have "F ?g ... = ?r" using reindex [OF inj2, of ?g] by simp finally show ?thesis. qed end subsection \Sum over a set of permutations (could generalize to iteration)\ lemma sum_over_permutations_insert: assumes fS: "finite S" and aS: "a \ S" shows "sum f {p. p permutes (insert a S)} = sum (\b. sum (\q. f (transpose a b \ q)) {p. p permutes S}) (insert a S)" proof - have *: "\f a b. (\(b, p). f (transpose a b \ p)) = f \ (\(b,p). transpose a b \ p)" by (simp add: fun_eq_iff) have **: "\P Q. {(a, b). a \ P \ b \ Q} = P \ Q" by blast show ?thesis unfolding * ** sum.cartesian_product permutes_insert proof (rule sum.reindex) let ?f = "(\(b, y). transpose a b \ y)" let ?P = "{p. p permutes S}" { fix b c p q assume b: "b \ insert a S" assume c: "c \ insert a S" assume p: "p permutes S" assume q: "q permutes S" assume eq: "transpose a b \ p = transpose a c \ q" from p q aS have pa: "p a = a" and qa: "q a = a" unfolding permutes_def by metis+ from eq have "(transpose a b \ p) a = (transpose a c \ q) a" by simp then have bc: "b = c" by (simp add: permutes_def pa qa o_def fun_upd_def id_def cong del: if_weak_cong split: if_split_asm) from eq[unfolded bc] have "(\p. transpose a c \ p) (transpose a c \ p) = (\p. transpose a c \ p) (transpose a c \ q)" by simp then have "p = q" unfolding o_assoc swap_id_idempotent by simp with bc have "b = c \ p = q" by blast } then show "inj_on ?f (insert a S \ ?P)" unfolding inj_on_def by clarify metis qed qed subsection \Constructing permutations from association lists\ definition list_permutes :: "('a \ 'a) list \ 'a set \ bool" where "list_permutes xs A \ set (map fst xs) \ A \ set (map snd xs) = set (map fst xs) \ distinct (map fst xs) \ distinct (map snd xs)" lemma list_permutesI [simp]: assumes "set (map fst xs) \ A" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)" shows "list_permutes xs A" proof - from assms(2,3) have "distinct (map snd xs)" by (intro card_distinct) (simp_all add: distinct_card del: set_map) with assms show ?thesis by (simp add: list_permutes_def) qed definition permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "permutation_of_list xs x = (case map_of xs x of None \ x | Some y \ y)" lemma permutation_of_list_Cons: "permutation_of_list ((x, y) # xs) x' = (if x = x' then y else permutation_of_list xs x')" by (simp add: permutation_of_list_def) fun inverse_permutation_of_list :: "('a \ 'a) list \ 'a \ 'a" where "inverse_permutation_of_list [] x = x" | "inverse_permutation_of_list ((y, x') # xs) x = (if x = x' then y else inverse_permutation_of_list xs x)" declare inverse_permutation_of_list.simps [simp del] lemma inj_on_map_of: assumes "distinct (map snd xs)" shows "inj_on (map_of xs) (set (map fst xs))" proof (rule inj_onI) fix x y assume xy: "x \ set (map fst xs)" "y \ set (map fst xs)" assume eq: "map_of xs x = map_of xs y" from xy obtain x' y' where x'y': "map_of xs x = Some x'" "map_of xs y = Some y'" by (cases "map_of xs x"; cases "map_of xs y") (simp_all add: map_of_eq_None_iff) moreover from x'y' have *: "(x, x') \ set xs" "(y, y') \ set xs" by (force dest: map_of_SomeD)+ moreover from * eq x'y' have "x' = y'" by simp ultimately show "x = y" using assms by (force simp: distinct_map dest: inj_onD[of _ _ "(x,x')" "(y,y')"]) qed lemma inj_on_the: "None \ A \ inj_on the A" by (auto simp: inj_on_def option.the_def split: option.splits) lemma inj_on_map_of': assumes "distinct (map snd xs)" shows "inj_on (the \ map_of xs) (set (map fst xs))" by (intro comp_inj_on inj_on_map_of assms inj_on_the) (force simp: eq_commute[of None] map_of_eq_None_iff) lemma image_map_of: assumes "distinct (map fst xs)" shows "map_of xs ` set (map fst xs) = Some ` set (map snd xs)" using assms by (auto simp: rev_image_eqI) lemma the_Some_image [simp]: "the ` Some ` A = A" by (subst image_image) simp lemma image_map_of': assumes "distinct (map fst xs)" shows "(the \ map_of xs) ` set (map fst xs) = set (map snd xs)" by (simp only: image_comp [symmetric] image_map_of assms the_Some_image) lemma permutation_of_list_permutes [simp]: assumes "list_permutes xs A" shows "permutation_of_list xs permutes A" (is "?f permutes _") proof (rule permutes_subset[OF bij_imp_permutes]) from assms show "set (map fst xs) \ A" by (simp add: list_permutes_def) from assms have "inj_on (the \ map_of xs) (set (map fst xs))" (is ?P) by (intro inj_on_map_of') (simp_all add: list_permutes_def) also have "?P \ inj_on ?f (set (map fst xs))" by (intro inj_on_cong) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) finally have "bij_betw ?f (set (map fst xs)) (?f ` set (map fst xs))" by (rule inj_on_imp_bij_betw) also from assms have "?f ` set (map fst xs) = (the \ map_of xs) ` set (map fst xs)" by (intro image_cong refl) (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits) also from assms have "\ = set (map fst xs)" by (subst image_map_of') (simp_all add: list_permutes_def) finally show "bij_betw ?f (set (map fst xs)) (set (map fst xs))" . qed (force simp: permutation_of_list_def dest!: map_of_SomeD split: option.splits)+ lemma eval_permutation_of_list [simp]: "permutation_of_list [] x = x" "x = x' \ permutation_of_list ((x',y)#xs) x = y" "x \ x' \ permutation_of_list ((x',y')#xs) x = permutation_of_list xs x" by (simp_all add: permutation_of_list_def) lemma eval_inverse_permutation_of_list [simp]: "inverse_permutation_of_list [] x = x" "x = x' \ inverse_permutation_of_list ((y,x')#xs) x = y" "x \ x' \ inverse_permutation_of_list ((y',x')#xs) x = inverse_permutation_of_list xs x" by (simp_all add: inverse_permutation_of_list.simps) lemma permutation_of_list_id: "x \ set (map fst xs) \ permutation_of_list xs x = x" by (induct xs) (auto simp: permutation_of_list_Cons) lemma permutation_of_list_unique': "distinct (map fst xs) \ (x, y) \ set xs \ permutation_of_list xs x = y" by (induct xs) (force simp: permutation_of_list_Cons)+ lemma permutation_of_list_unique: "list_permutes xs A \ (x, y) \ set xs \ permutation_of_list xs x = y" by (intro permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_id: "x \ set (map snd xs) \ inverse_permutation_of_list xs x = x" by (induct xs) auto lemma inverse_permutation_of_list_unique': "distinct (map snd xs) \ (x, y) \ set xs \ inverse_permutation_of_list xs y = x" by (induct xs) (force simp: inverse_permutation_of_list.simps(2))+ lemma inverse_permutation_of_list_unique: "list_permutes xs A \ (x,y) \ set xs \ inverse_permutation_of_list xs y = x" by (intro inverse_permutation_of_list_unique') (simp_all add: list_permutes_def) lemma inverse_permutation_of_list_correct: fixes A :: "'a set" assumes "list_permutes xs A" shows "inverse_permutation_of_list xs = inv (permutation_of_list xs)" proof (rule ext, rule sym, subst permutes_inv_eq) from assms show "permutation_of_list xs permutes A" by simp show "permutation_of_list xs (inverse_permutation_of_list xs x) = x" for x proof (cases "x \ set (map snd xs)") case True then obtain y where "(y, x) \ set xs" by auto with assms show ?thesis by (simp add: inverse_permutation_of_list_unique permutation_of_list_unique) next case False with assms show ?thesis by (auto simp: list_permutes_def inverse_permutation_of_list_id permutation_of_list_id) qed 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,3973 +1,4002 @@ (* 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 *) 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 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 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)" 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 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 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 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 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 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 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") = \fn phi => 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") = \fn phi => 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") = \fn phi => 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") = \fn phi => 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 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]) 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 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 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 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) 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) lemma size_empty [simp]: "size {#} = 0" by (simp add: size_multiset_overloaded_def) lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)" by (simp add: size_multiset_eq) lemma size_single: "size {#b#} = 1" 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 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) lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)" 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) 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) 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 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 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]) 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) 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 lemma multiset_cases [cases type]: 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) 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 lemma multiset_partition: "M = {#x \# M. P x#} + {#x \# M. \ P x#}" 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) qed lemma size_1_singleton_mset: "size M = 1 \ \a. M = {#a#}" by (cases M) 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 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 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 \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) 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_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_funpow) with * show ?thesis by (simp add: fold_mset_def del: count_add_mset) simp qed qed corollary fold_mset_single: "fold_mset f s {#x#} = f x s" by simp 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 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" apply (unfold surj_def) apply (rule allI) apply (rule_tac M = y in multiset_induct) apply auto apply (rule_tac x = "x # xa" in exI) apply auto done 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 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: - assumes f: "\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x" - and equiv: "mset xs = mset ys" - shows "List.fold f xs = List.fold f ys" - using f equiv [symmetric] -proof (induct xs arbitrary: ys) + \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)" + 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" + 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" + moreover from * have \x \ set ys\ by simp - ultimately have "List.fold f ys = List.fold f (remove1 x ys) \ f x" + 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.hyps) - ultimately show ?case by simp + 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.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) 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) 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 and repeat 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) lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" by (induct n, simp_all) 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 filter_eq_replicate_mset: "{#y \# D. y = x#} = replicate_mset (count D x) x" by (induct D) simp_all 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 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 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 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\ subsubsection \Well-foundedness\ 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>+" 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 not_less_empty [iff]: "(M, {#}) \ mult1 r" by (simp add: mult1_def) 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 theorem wf_mult1: "wf r \ wf (mult1 r)" by (rule acc_wfI) (rule all_accessible) theorem wf_mult: "wf r \ wf (mult r)" unfolding mult_def by (rule wf_trancl) (rule wf_mult1) 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 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 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 subsection \The multiset extension is cancellative for multiset union\ lemma mult_cancel: assumes "trans s" and "irrefl s" shows "(X + Z, Y + Z) \ mult s \ (X, Y) \ mult s" (is "?L \ ?R") proof assume ?L thus ?R 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] by (auto simp: add.commute[of _ "{#_#}"] add.assoc intro: add(1)) next case 2 then obtain y where "y \ set_mset Y2" "(z, y) \ s" using *(4) \irrefl s\ by (auto simp: irrefl_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 by (force simp: add.commute[of "{#_#}"] add.assoc[symmetric] intro: add(1)) 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 lemmas mult_cancel_add_mset = mult_cancel[of _ _ "{#_#}", unfolded union_mset_add_mset_right add.comm_neutral] lemma mult_cancel_max: assumes "trans s" and "irrefl s" shows "(X, Y) \ mult s \ (X - X \# Y, Y - X \# Y) \ mult s" (is "?L \ ?R") proof - have "X - X \# Y + X \# Y = X" "Y - X \# Y + X \# Y = Y" by (auto simp flip: count_inject) thus ?thesis using mult_cancel[OF assms, of "X - X \# Y" "X \# Y" "Y - X \# Y"] by auto qed 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\ and \multeqp\. \ definition multp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp 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 :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multeqp 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_iff: assumes "irrefl R" and "trans R" and [simp]: "\x y. P x y \ (x, y) \ R" shows "multp 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_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_implies_one_step[OF assms(2), of "N - M \# N" "M - M \# N"] mult_cancel_max[OF assms(2,1), of "N" "M"] * by (auto simp: multp_def) qed qed lemma multeqp_iff: assumes "irrefl R" and "trans R" and "\x y. P x y \ (x, y) \ R" shows "multeqp 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 P N M \ multp P N M \ N = M" by (auto simp: multeqp_def multp_def Let_def in_diff_count) thus ?thesis using multp_iff[OF assms] by simp qed subsubsection \Partial-order properties\ lemma (in preorder) mult1_lessE: assumes "(N, M) \ mult1 {(a, b). a < b}" obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "a \# K" "\b. b \# K \ b < a" proof - from assms obtain a M0 K where "M = add_mset a M0" "N = M0 + K" and *: "b \# K \ b < a" for b by (blast elim: mult1E) moreover from * [of a] have "a \# K" by auto ultimately show thesis by (auto intro: that) qed instantiation multiset :: (preorder) order begin definition less_multiset :: "'a multiset \ 'a multiset \ bool" where "M' < M \ (M', M) \ mult {(x', x). x' < x}" definition less_eq_multiset :: "'a multiset \ 'a multiset \ bool" where "less_eq_multiset M' M \ M' < M \ M' = M" instance proof - have irrefl: "\ M < M" for M :: "'a multiset" proof assume "M < M" then have MM: "(M, M) \ mult {(x, y). x < y}" by (simp add: less_multiset_def) have "trans {(x'::'a, x). x' < x}" by (metis (mono_tags, lifting) case_prodD case_prodI less_trans mem_Collect_eq transI) moreover note MM ultimately have "\I J K. M = I + J \ M = I + K \ J \ {#} \ (\k\set_mset K. \j\set_mset J. (k, j) \ {(x, y). x < y})" by (rule mult_implies_one_step) 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) \ {(x, y). x < y})" by blast then have *: "K \ {#}" and **: "\k\set_mset K. \j\set_mset K. k < j" by auto have "finite (set_mset K)" by simp moreover note ** ultimately have "set_mset K = {}" by (induct rule: finite_induct) (auto intro: order_less_trans) with * show False by simp qed have trans: "K < M \ M < N \ K < N" for K M N :: "'a multiset" unfolding less_multiset_def mult_def by (blast intro: trancl_trans) show "OFCLASS('a multiset, order_class)" by standard (auto simp add: less_eq_multiset_def irrefl dest: trans) qed end lemma mset_le_irrefl [elim!]: fixes M :: "'a::preorder multiset" shows "M < M \ R" by simp 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 mult_def) apply (erule trancl_induct) apply (blast intro: mult1_union) apply (blast intro: mult1_union trancl_trans) done 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 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) 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 (\<^type_name>\multiset\, [T]); fun mk_mset T [] = Const (\<^const_abbrev>\empty_mset\, msetT T) | mk_mset T [x] = Const (\<^const_name>\add_mset\, T --> msetT T --> msetT T) $ x $ Const (\<^const_abbrev>\empty_mset\, msetT T) | mk_mset T (x :: xs) = Const (\<^const_name>\plus\, msetT T --> msetT T --> msetT T) $ 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 (\<^const_name>\zero_class.zero\, T) | ts => foldl1 (fn (s, t) => Const (\<^const_name>\add_mset\, elem_T --> T --> T) $ 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 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 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) 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) 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 "ordLeq3 (card_of (set_mset X)) natLeq" for X by transfer (auto intro!: ordLess_imp_ordLeq 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 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 show "pred_mset P = (\x. Ball (set_mset x) P)" for P by (simp add: fun_eq_iff pred_mset_iff) qed 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 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 hide_const (open) wcount end