wenzelm@10249: (* Title: HOL/Library/Multiset.thy wenzelm@10249: ID: $Id$ paulson@15072: Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker wenzelm@10249: *) wenzelm@10249: wenzelm@14706: header {* Multisets *} wenzelm@10249: nipkow@15131: theory Multiset haftmann@27487: imports Plain "~~/src/HOL/List" nipkow@15131: begin wenzelm@10249: wenzelm@10249: subsection {* The type of multisets *} wenzelm@10249: nipkow@25162: typedef 'a multiset = "{f::'a => nat. finite {x . f x > 0}}" wenzelm@10249: proof nipkow@11464: show "(\x. 0::nat) \ ?multiset" by simp wenzelm@10249: qed wenzelm@10249: wenzelm@10249: lemmas multiset_typedef [simp] = wenzelm@10277: Abs_multiset_inverse Rep_multiset_inverse Rep_multiset wenzelm@10277: and [simp] = Rep_multiset_inject [symmetric] wenzelm@10249: haftmann@28708: definition Mempty :: "'a multiset" ("{#}") where haftmann@28708: [code del]: "{#} = Abs_multiset (\a. 0)" wenzelm@10249: haftmann@28708: definition single :: "'a => 'a multiset" where haftmann@28708: [code del]: "single a = Abs_multiset (\b. if b = a then 1 else 0)" wenzelm@10249: haftmann@28708: definition count :: "'a multiset => 'a => nat" where wenzelm@19086: "count = Rep_multiset" wenzelm@10249: haftmann@28708: definition MCollect :: "'a multiset => ('a => bool) => 'a multiset" where wenzelm@19086: "MCollect M P = Abs_multiset (\x. if P x then Rep_multiset M x else 0)" wenzelm@19086: haftmann@28708: abbreviation Melem :: "'a => 'a multiset => bool" ("(_/ :# _)" [50, 51] 50) where kleing@25610: "a :# M == 0 < count M a" kleing@25610: wenzelm@26145: notation (xsymbols) wenzelm@26145: Melem (infix "\#" 50) wenzelm@10249: wenzelm@10249: syntax nipkow@26033: "_MCollect" :: "pttrn => 'a multiset => bool => 'a multiset" ("(1{# _ :# _./ _#})") wenzelm@10249: translations nipkow@26033: "{#x :# M. P#}" == "CONST MCollect M (\x. P)" wenzelm@10249: haftmann@28708: definition set_of :: "'a multiset => 'a set" where wenzelm@19086: "set_of M = {x. x :# M}" wenzelm@10249: haftmann@25571: instantiation multiset :: (type) "{plus, minus, zero, size}" haftmann@25571: begin haftmann@25571: haftmann@28708: definition union_def [code del]: wenzelm@26145: "M + N = Abs_multiset (\a. Rep_multiset M a + Rep_multiset N a)" haftmann@25571: haftmann@28708: definition diff_def [code del]: haftmann@28708: "M - N = Abs_multiset (\a. Rep_multiset M a - Rep_multiset N a)" haftmann@25571: haftmann@28708: definition Zero_multiset_def [simp]: haftmann@28708: "0 = {#}" haftmann@25571: haftmann@28708: definition size_def: haftmann@28708: "size M = setsum (count M) (set_of M)" haftmann@25571: haftmann@25571: instance .. haftmann@25571: haftmann@25571: end wenzelm@10249: wenzelm@19086: definition wenzelm@21404: multiset_inter :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "#\" 70) where wenzelm@19086: "multiset_inter A B = A - (A - B)" kleing@15869: wenzelm@26145: text {* Multiset Enumeration *} wenzelm@26145: syntax wenzelm@26176: "_multiset" :: "args => 'a multiset" ("{#(_)#}") nipkow@25507: translations nipkow@25507: "{#x, xs#}" == "{#x#} + {#xs#}" nipkow@25507: "{#x#}" == "CONST single x" nipkow@25507: wenzelm@10249: wenzelm@10249: text {* wenzelm@10249: \medskip Preservation of the representing set @{term multiset}. wenzelm@10249: *} wenzelm@10249: nipkow@26016: lemma const0_in_multiset: "(\a. 0) \ multiset" nipkow@26178: by (simp add: multiset_def) wenzelm@10249: nipkow@26016: lemma only1_in_multiset: "(\b. if b = a then 1 else 0) \ multiset" nipkow@26178: by (simp add: multiset_def) wenzelm@10249: nipkow@26016: lemma union_preserves_multiset: nipkow@26178: "M \ multiset ==> N \ multiset ==> (\a. M a + N a) \ multiset" nipkow@26178: apply (simp add: multiset_def) nipkow@26178: apply (drule (1) finite_UnI) nipkow@26178: apply (simp del: finite_Un add: Un_def) nipkow@26178: done wenzelm@10249: nipkow@26016: lemma diff_preserves_multiset: nipkow@26178: "M \ multiset ==> (\a. M a - N a) \ multiset" nipkow@26178: apply (simp add: multiset_def) nipkow@26178: apply (rule finite_subset) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: nipkow@26016: lemma MCollect_preserves_multiset: nipkow@26178: "M \ multiset ==> (\x. if P x then M x else 0) \ multiset" nipkow@26178: apply (simp add: multiset_def) nipkow@26178: apply (rule finite_subset, auto) nipkow@26178: done wenzelm@10249: nipkow@26016: lemmas in_multiset = const0_in_multiset only1_in_multiset nipkow@26016: union_preserves_multiset diff_preserves_multiset MCollect_preserves_multiset nipkow@26016: wenzelm@26145: nipkow@26016: subsection {* Algebraic properties *} wenzelm@10249: wenzelm@10249: subsubsection {* Union *} wenzelm@10249: wenzelm@17161: lemma union_empty [simp]: "M + {#} = M \ {#} + M = M" nipkow@26178: by (simp add: union_def Mempty_def in_multiset) wenzelm@10249: wenzelm@17161: lemma union_commute: "M + N = N + (M::'a multiset)" nipkow@26178: by (simp add: union_def add_ac in_multiset) wenzelm@17161: wenzelm@17161: lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" nipkow@26178: by (simp add: union_def add_ac in_multiset) wenzelm@10249: wenzelm@17161: lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" wenzelm@17161: proof - nipkow@26178: have "M + (N + K) = (N + K) + M" by (rule union_commute) nipkow@26178: also have "\ = N + (K + M)" by (rule union_assoc) nipkow@26178: also have "K + M = M + K" by (rule union_commute) wenzelm@17161: finally show ?thesis . wenzelm@17161: qed wenzelm@10249: wenzelm@17161: lemmas union_ac = union_assoc union_commute union_lcomm wenzelm@10249: obua@14738: instance multiset :: (type) comm_monoid_add wenzelm@17200: proof obua@14722: fix a b c :: "'a multiset" obua@14722: show "(a + b) + c = a + (b + c)" by (rule union_assoc) obua@14722: show "a + b = b + a" by (rule union_commute) obua@14722: show "0 + a = a" by simp obua@14722: qed wenzelm@10277: wenzelm@10249: wenzelm@10249: subsubsection {* Difference *} wenzelm@10249: wenzelm@17161: lemma diff_empty [simp]: "M - {#} = M \ {#} - M = {#}" nipkow@26178: by (simp add: Mempty_def diff_def in_multiset) wenzelm@10249: wenzelm@17161: lemma diff_union_inverse2 [simp]: "M + {#a#} - {#a#} = M" nipkow@26178: by (simp add: union_def diff_def in_multiset) wenzelm@10249: bulwahn@26143: lemma diff_cancel: "A - A = {#}" nipkow@26178: by (simp add: diff_def Mempty_def) bulwahn@26143: wenzelm@10249: wenzelm@10249: subsubsection {* Count of elements *} wenzelm@10249: wenzelm@17161: lemma count_empty [simp]: "count {#} a = 0" nipkow@26178: by (simp add: count_def Mempty_def in_multiset) wenzelm@10249: wenzelm@17161: lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)" nipkow@26178: by (simp add: count_def single_def in_multiset) wenzelm@10249: wenzelm@17161: lemma count_union [simp]: "count (M + N) a = count M a + count N a" nipkow@26178: by (simp add: count_def union_def in_multiset) wenzelm@10249: wenzelm@17161: lemma count_diff [simp]: "count (M - N) a = count M a - count N a" nipkow@26178: by (simp add: count_def diff_def in_multiset) nipkow@26016: nipkow@26016: lemma count_MCollect [simp]: nipkow@26178: "count {# x:#M. P x #} a = (if P a then count M a else 0)" nipkow@26178: by (simp add: count_def MCollect_def in_multiset) wenzelm@10249: wenzelm@10249: wenzelm@10249: subsubsection {* Set of elements *} wenzelm@10249: wenzelm@17161: lemma set_of_empty [simp]: "set_of {#} = {}" nipkow@26178: by (simp add: set_of_def) wenzelm@10249: wenzelm@17161: lemma set_of_single [simp]: "set_of {#b#} = {b}" nipkow@26178: by (simp add: set_of_def) wenzelm@10249: wenzelm@17161: lemma set_of_union [simp]: "set_of (M + N) = set_of M \ set_of N" nipkow@26178: by (auto simp add: set_of_def) wenzelm@10249: wenzelm@17161: lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})" berghofe@26818: by (auto simp: set_of_def Mempty_def in_multiset count_def expand_fun_eq [where f="Rep_multiset M"]) wenzelm@10249: wenzelm@17161: lemma mem_set_of_iff [simp]: "(x \ set_of M) = (x :# M)" nipkow@26178: by (auto simp add: set_of_def) nipkow@26016: nipkow@26033: lemma set_of_MCollect [simp]: "set_of {# x:#M. P x #} = set_of M \ {x. P x}" nipkow@26178: by (auto simp add: set_of_def) wenzelm@10249: wenzelm@10249: wenzelm@10249: subsubsection {* Size *} wenzelm@10249: haftmann@28708: lemma size_empty [simp]: "size {#} = 0" nipkow@26178: by (simp add: size_def) wenzelm@10249: haftmann@28708: lemma size_single [simp]: "size {#b#} = 1" nipkow@26178: by (simp add: size_def) wenzelm@10249: wenzelm@17161: lemma finite_set_of [iff]: "finite (set_of M)" nipkow@26178: using Rep_multiset [of M] by (simp add: multiset_def set_of_def count_def) wenzelm@10249: wenzelm@17161: lemma setsum_count_Int: nipkow@26178: "finite A ==> setsum (count N) (A \ set_of N) = setsum (count N) A" nipkow@26178: apply (induct rule: finite_induct) nipkow@26178: apply simp nipkow@26178: apply (simp add: Int_insert_left set_of_def) nipkow@26178: done wenzelm@10249: haftmann@28708: lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" nipkow@26178: apply (unfold size_def) nipkow@26178: apply (subgoal_tac "count (M + N) = (\a. count M a + count N a)") nipkow@26178: prefer 2 nipkow@26178: apply (rule ext, simp) nipkow@26178: apply (simp (no_asm_simp) add: setsum_Un_nat setsum_addf setsum_count_Int) nipkow@26178: apply (subst Int_commute) nipkow@26178: apply (simp (no_asm_simp) add: setsum_count_Int) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" nipkow@26178: apply (unfold size_def Mempty_def count_def, auto simp: in_multiset) nipkow@26178: apply (simp add: set_of_def count_def in_multiset expand_fun_eq) nipkow@26178: done nipkow@26016: nipkow@26016: lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" nipkow@26178: by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) wenzelm@10249: wenzelm@17161: lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \a. a :# M" nipkow@26178: apply (unfold size_def) nipkow@26178: apply (drule setsum_SucD) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: wenzelm@26145: wenzelm@10249: subsubsection {* Equality of multisets *} wenzelm@10249: wenzelm@17161: lemma multiset_eq_conv_count_eq: "(M = N) = (\a. count M a = count N a)" nipkow@26178: by (simp add: count_def expand_fun_eq) wenzelm@10249: wenzelm@17161: lemma single_not_empty [simp]: "{#a#} \ {#} \ {#} \ {#a#}" nipkow@26178: by (simp add: single_def Mempty_def in_multiset expand_fun_eq) wenzelm@10249: wenzelm@17161: lemma single_eq_single [simp]: "({#a#} = {#b#}) = (a = b)" nipkow@26178: by (auto simp add: single_def in_multiset expand_fun_eq) wenzelm@10249: wenzelm@17161: lemma union_eq_empty [iff]: "(M + N = {#}) = (M = {#} \ N = {#})" nipkow@26178: by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) wenzelm@10249: wenzelm@17161: lemma empty_eq_union [iff]: "({#} = M + N) = (M = {#} \ N = {#})" nipkow@26178: by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) wenzelm@10249: wenzelm@17161: lemma union_right_cancel [simp]: "(M + K = N + K) = (M = (N::'a multiset))" nipkow@26178: by (simp add: union_def in_multiset expand_fun_eq) wenzelm@10249: wenzelm@17161: lemma union_left_cancel [simp]: "(K + M = K + N) = (M = (N::'a multiset))" nipkow@26178: by (simp add: union_def in_multiset expand_fun_eq) wenzelm@10249: wenzelm@17161: lemma union_is_single: nipkow@26178: "(M + N = {#a#}) = (M = {#a#} \ N={#} \ M = {#} \ N = {#a#})" nipkow@26178: apply (simp add: Mempty_def single_def union_def in_multiset add_is_1 expand_fun_eq) nipkow@26178: apply blast nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma single_is_union: nipkow@26178: "({#a#} = M + N) \ ({#a#} = M \ N = {#} \ M = {#} \ {#a#} = N)" nipkow@26178: apply (unfold Mempty_def single_def union_def) nipkow@26178: apply (simp add: add_is_1 one_is_add in_multiset expand_fun_eq) nipkow@26178: apply (blast dest: sym) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma add_eq_conv_diff: wenzelm@10249: "(M + {#a#} = N + {#b#}) = paulson@15072: (M = N \ a = b \ M = N - {#a#} + {#b#} \ N = M - {#b#} + {#a#})" nipkow@26178: using [[simproc del: neq]] nipkow@26178: apply (unfold single_def union_def diff_def) nipkow@26178: apply (simp (no_asm) add: in_multiset expand_fun_eq) nipkow@26178: apply (rule conjI, force, safe, simp_all) nipkow@26178: apply (simp add: eq_sym_conv) nipkow@26178: done wenzelm@10249: kleing@15869: declare Rep_multiset_inject [symmetric, simp del] kleing@15869: nipkow@23611: instance multiset :: (type) cancel_ab_semigroup_add nipkow@23611: proof nipkow@23611: fix a b c :: "'a multiset" nipkow@23611: show "a + b = a + c \ b = c" by simp nipkow@23611: qed kleing@15869: kleing@25610: lemma insert_DiffM: kleing@25610: "x \# M \ {#x#} + (M - {#x#}) = M" nipkow@26178: by (clarsimp simp: multiset_eq_conv_count_eq) kleing@25610: kleing@25610: lemma insert_DiffM2[simp]: kleing@25610: "x \# M \ M - {#x#} + {#x#} = M" nipkow@26178: by (clarsimp simp: multiset_eq_conv_count_eq) kleing@25610: kleing@25610: lemma multi_union_self_other_eq: kleing@25610: "(A::'a multiset) + X = A + Y \ X = Y" nipkow@26178: by (induct A arbitrary: X Y) auto kleing@25610: kleing@25610: lemma multi_self_add_other_not_self[simp]: "(A = A + {#x#}) = False" nipkow@26178: by (metis single_not_empty union_empty union_left_cancel) kleing@25610: kleing@25610: lemma insert_noteq_member: kleing@25610: assumes BC: "B + {#b#} = C + {#c#}" kleing@25610: and bnotc: "b \ c" kleing@25610: shows "c \# B" kleing@25610: proof - kleing@25610: have "c \# C + {#c#}" by simp kleing@25610: have nc: "\ c \# {#b#}" using bnotc by simp wenzelm@26145: then have "c \# B + {#b#}" using BC by simp wenzelm@26145: then show "c \# B" using nc by simp kleing@25610: qed kleing@25610: kleing@25610: nipkow@26016: lemma add_eq_conv_ex: nipkow@26016: "(M + {#a#} = N + {#b#}) = nipkow@26016: (M = N \ a = b \ (\K. M = K + {#b#} \ N = K + {#a#}))" nipkow@26178: by (auto simp add: add_eq_conv_diff) nipkow@26016: nipkow@26016: nipkow@26016: lemma empty_multiset_count: nipkow@26016: "(\x. count A x = 0) = (A = {#})" nipkow@26178: by (metis count_empty multiset_eq_conv_count_eq) nipkow@26016: nipkow@26016: kleing@15869: subsubsection {* Intersection *} kleing@15869: kleing@15869: lemma multiset_inter_count: nipkow@26178: "count (A #\ B) x = min (count A x) (count B x)" nipkow@26178: by (simp add: multiset_inter_def min_def) kleing@15869: kleing@15869: lemma multiset_inter_commute: "A #\ B = B #\ A" nipkow@26178: by (simp add: multiset_eq_conv_count_eq multiset_inter_count haftmann@21214: min_max.inf_commute) kleing@15869: kleing@15869: lemma multiset_inter_assoc: "A #\ (B #\ C) = A #\ B #\ C" nipkow@26178: by (simp add: multiset_eq_conv_count_eq multiset_inter_count haftmann@21214: min_max.inf_assoc) kleing@15869: kleing@15869: lemma multiset_inter_left_commute: "A #\ (B #\ C) = B #\ (A #\ C)" nipkow@26178: by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def) kleing@15869: wenzelm@17161: lemmas multiset_inter_ac = wenzelm@17161: multiset_inter_commute wenzelm@17161: multiset_inter_assoc wenzelm@17161: multiset_inter_left_commute kleing@15869: bulwahn@26143: lemma multiset_inter_single: "a \ b \ {#a#} #\ {#b#} = {#}" nipkow@26178: by (simp add: multiset_eq_conv_count_eq multiset_inter_count) bulwahn@26143: kleing@15869: lemma multiset_union_diff_commute: "B #\ C = {#} \ A + B - C = A - C + B" nipkow@26178: apply (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def wenzelm@17161: split: split_if_asm) nipkow@26178: apply clarsimp nipkow@26178: apply (erule_tac x = a in allE) nipkow@26178: apply auto nipkow@26178: done kleing@15869: wenzelm@10249: nipkow@26016: subsubsection {* Comprehension (filter) *} nipkow@26016: haftmann@28708: lemma MCollect_empty [simp]: "MCollect {#} P = {#}" nipkow@26178: by (simp add: MCollect_def Mempty_def Abs_multiset_inject wenzelm@26145: in_multiset expand_fun_eq) nipkow@26016: haftmann@28708: lemma MCollect_single [simp]: nipkow@26178: "MCollect {#x#} P = (if P x then {#x#} else {#})" nipkow@26178: by (simp add: MCollect_def Mempty_def single_def Abs_multiset_inject wenzelm@26145: in_multiset expand_fun_eq) nipkow@26016: haftmann@28708: lemma MCollect_union [simp]: nipkow@26016: "MCollect (M+N) f = MCollect M f + MCollect N f" nipkow@26178: by (simp add: MCollect_def union_def Abs_multiset_inject wenzelm@26145: in_multiset expand_fun_eq) nipkow@26016: nipkow@26016: nipkow@26016: subsection {* Induction and case splits *} wenzelm@10249: wenzelm@10249: lemma setsum_decr: wenzelm@11701: "finite F ==> (0::nat) < f a ==> paulson@15072: setsum (f (a := f a - 1)) F = (if a\F then setsum f F - 1 else setsum f F)" nipkow@26178: apply (induct rule: finite_induct) nipkow@26178: apply auto nipkow@26178: apply (drule_tac a = a in mk_disjoint_insert, auto) nipkow@26178: done wenzelm@10249: wenzelm@10313: lemma rep_multiset_induct_aux: nipkow@26178: assumes 1: "P (\a. (0::nat))" nipkow@26178: and 2: "!!f b. f \ multiset ==> P f ==> P (f (b := f b + 1))" nipkow@26178: shows "\f. f \ multiset --> setsum f {x. f x \ 0} = n --> P f" nipkow@26178: apply (unfold multiset_def) nipkow@26178: apply (induct_tac n, simp, clarify) nipkow@26178: apply (subgoal_tac "f = (\a.0)") nipkow@26178: apply simp nipkow@26178: apply (rule 1) nipkow@26178: apply (rule ext, force, clarify) nipkow@26178: apply (frule setsum_SucD, clarify) nipkow@26178: apply (rename_tac a) nipkow@26178: apply (subgoal_tac "finite {x. (f (a := f a - 1)) x > 0}") nipkow@26178: prefer 2 nipkow@26178: apply (rule finite_subset) nipkow@26178: prefer 2 nipkow@26178: apply assumption nipkow@26178: apply simp nipkow@26178: apply blast nipkow@26178: apply (subgoal_tac "f = (f (a := f a - 1))(a := (f (a := f a - 1)) a + 1)") nipkow@26178: prefer 2 nipkow@26178: apply (rule ext) nipkow@26178: apply (simp (no_asm_simp)) nipkow@26178: apply (erule ssubst, rule 2 [unfolded multiset_def], blast) nipkow@26178: apply (erule allE, erule impE, erule_tac [2] mp, blast) nipkow@26178: apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply One_nat_def) nipkow@26178: apply (subgoal_tac "{x. x \ a --> f x \ 0} = {x. f x \ 0}") nipkow@26178: prefer 2 nipkow@26178: apply blast nipkow@26178: apply (subgoal_tac "{x. x \ a \ f x \ 0} = {x. f x \ 0} - {a}") nipkow@26178: prefer 2 nipkow@26178: apply blast nipkow@26178: apply (simp add: le_imp_diff_is_add setsum_diff1_nat cong: conj_cong) nipkow@26178: done wenzelm@10249: wenzelm@10313: theorem rep_multiset_induct: nipkow@11464: "f \ multiset ==> P (\a. 0) ==> wenzelm@11701: (!!f b. f \ multiset ==> P f ==> P (f (b := f b + 1))) ==> P f" nipkow@26178: using rep_multiset_induct_aux by blast wenzelm@10249: wenzelm@18258: theorem multiset_induct [case_names empty add, induct type: multiset]: nipkow@26178: assumes empty: "P {#}" nipkow@26178: and add: "!!M x. P M ==> P (M + {#x#})" nipkow@26178: shows "P M" wenzelm@10249: proof - wenzelm@10249: note defns = union_def single_def Mempty_def wenzelm@10249: show ?thesis wenzelm@10249: apply (rule Rep_multiset_inverse [THEN subst]) wenzelm@10313: apply (rule Rep_multiset [THEN rep_multiset_induct]) wenzelm@18258: apply (rule empty [unfolded defns]) paulson@15072: apply (subgoal_tac "f(b := f b + 1) = (\a. f a + (if a=b then 1 else 0))") wenzelm@10249: prefer 2 wenzelm@10249: apply (simp add: expand_fun_eq) wenzelm@10249: apply (erule ssubst) wenzelm@17200: apply (erule Abs_multiset_inverse [THEN subst]) nipkow@26016: apply (drule add [unfolded defns, simplified]) nipkow@26016: apply(simp add:in_multiset) wenzelm@10249: done wenzelm@10249: qed wenzelm@10249: kleing@25610: lemma multi_nonempty_split: "M \ {#} \ \A a. M = A + {#a#}" nipkow@26178: by (induct M) auto kleing@25610: kleing@25610: lemma multiset_cases [cases type, case_names empty add]: nipkow@26178: assumes em: "M = {#} \ P" nipkow@26178: assumes add: "\N x. M = N + {#x#} \ P" nipkow@26178: shows "P" kleing@25610: proof (cases "M = {#}") wenzelm@26145: assume "M = {#}" then show ?thesis using em by simp kleing@25610: next kleing@25610: assume "M \ {#}" kleing@25610: then obtain M' m where "M = M' + {#m#}" kleing@25610: by (blast dest: multi_nonempty_split) wenzelm@26145: then show ?thesis using add by simp kleing@25610: qed kleing@25610: kleing@25610: lemma multi_member_split: "x \# M \ \A. M = A + {#x#}" nipkow@26178: apply (cases M) nipkow@26178: apply simp nipkow@26178: apply (rule_tac x="M - {#x#}" in exI, simp) nipkow@26178: done kleing@25610: nipkow@26033: lemma multiset_partition: "M = {# x:#M. P x #} + {# x:#M. \ P x #}" nipkow@26178: apply (subst multiset_eq_conv_count_eq) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: kleing@15869: declare multiset_typedef [simp del] wenzelm@10249: kleing@25610: lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" nipkow@26178: by (cases "B = {#}") (auto dest: multi_member_split) wenzelm@26145: wenzelm@17161: nipkow@26016: subsection {* Orderings *} wenzelm@10249: wenzelm@10249: subsubsection {* Well-foundedness *} wenzelm@10249: haftmann@28708: definition mult1 :: "('a \ 'a) set => ('a multiset \ 'a multiset) set" where haftmann@28708: [code del]: "mult1 r = {(N, M). \a M0 K. M = M0 + {#a#} \ N = M0 + K \ berghofe@23751: (\b. b :# K --> (b, a) \ r)}" wenzelm@10249: haftmann@28708: definition mult :: "('a \ 'a) set => ('a multiset \ 'a multiset) set" where berghofe@23751: "mult r = (mult1 r)\<^sup>+" wenzelm@10249: berghofe@23751: lemma not_less_empty [iff]: "(M, {#}) \ mult1 r" nipkow@26178: by (simp add: mult1_def) wenzelm@10249: berghofe@23751: lemma less_add: "(N, M0 + {#a#}) \ mult1 r ==> berghofe@23751: (\M. (M, M0) \ mult1 r \ N = M + {#a#}) \ berghofe@23751: (\K. (\b. b :# K --> (b, a) \ r) \ N = M0 + K)" wenzelm@19582: (is "_ \ ?case1 (mult1 r) \ ?case2") wenzelm@10249: proof (unfold mult1_def) berghofe@23751: let ?r = "\K a. \b. b :# K --> (b, a) \ r" nipkow@11464: let ?R = "\N M. \a M0 K. M = M0 + {#a#} \ N = M0 + K \ ?r K a" berghofe@23751: let ?case1 = "?case1 {(N, M). ?R N M}" wenzelm@10249: berghofe@23751: assume "(N, M0 + {#a#}) \ {(N, M). ?R N M}" wenzelm@18258: then have "\a' M0' K. nipkow@11464: M0 + {#a#} = M0' + {#a'#} \ N = M0' + K \ ?r K a'" by simp wenzelm@18258: then show "?case1 \ ?case2" wenzelm@10249: proof (elim exE conjE) wenzelm@10249: fix a' M0' K wenzelm@10249: assume N: "N = M0' + K" and r: "?r K a'" wenzelm@10249: assume "M0 + {#a#} = M0' + {#a'#}" wenzelm@18258: then have "M0 = M0' \ a = a' \ nipkow@11464: (\K'. M0 = K' + {#a'#} \ M0' = K' + {#a#})" wenzelm@10249: by (simp only: add_eq_conv_ex) wenzelm@18258: then show ?thesis wenzelm@10249: proof (elim disjE conjE exE) wenzelm@10249: assume "M0 = M0'" "a = a'" nipkow@11464: with N r have "?r K a \ N = M0 + K" by simp wenzelm@18258: then have ?case2 .. then show ?thesis .. wenzelm@10249: next wenzelm@10249: fix K' wenzelm@10249: assume "M0' = K' + {#a#}" wenzelm@10249: with N have n: "N = K' + K + {#a#}" by (simp add: union_ac) wenzelm@10249: wenzelm@10249: assume "M0 = K' + {#a'#}" wenzelm@10249: with r have "?R (K' + K) M0" by blast wenzelm@18258: with n have ?case1 by simp then show ?thesis .. wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: berghofe@23751: lemma all_accessible: "wf r ==> \M. M \ acc (mult1 r)" wenzelm@10249: proof wenzelm@10249: let ?R = "mult1 r" wenzelm@10249: let ?W = "acc ?R" wenzelm@10249: { wenzelm@10249: fix M M0 a berghofe@23751: assume M0: "M0 \ ?W" berghofe@23751: and wf_hyp: "!!b. (b, a) \ r ==> (\M \ ?W. M + {#b#} \ ?W)" berghofe@23751: and acc_hyp: "\M. (M, M0) \ ?R --> M + {#a#} \ ?W" berghofe@23751: have "M0 + {#a#} \ ?W" berghofe@23751: proof (rule accI [of "M0 + {#a#}"]) wenzelm@10249: fix N berghofe@23751: assume "(N, M0 + {#a#}) \ ?R" berghofe@23751: then have "((\M. (M, M0) \ ?R \ N = M + {#a#}) \ berghofe@23751: (\K. (\b. b :# K --> (b, a) \ r) \ N = M0 + K))" wenzelm@10249: by (rule less_add) berghofe@23751: then show "N \ ?W" wenzelm@10249: proof (elim exE disjE conjE) berghofe@23751: fix M assume "(M, M0) \ ?R" and N: "N = M + {#a#}" berghofe@23751: from acc_hyp have "(M, M0) \ ?R --> M + {#a#} \ ?W" .. berghofe@23751: from this and `(M, M0) \ ?R` have "M + {#a#} \ ?W" .. berghofe@23751: then show "N \ ?W" by (simp only: N) wenzelm@10249: next wenzelm@10249: fix K wenzelm@10249: assume N: "N = M0 + K" berghofe@23751: assume "\b. b :# K --> (b, a) \ r" berghofe@23751: then have "M0 + K \ ?W" wenzelm@10249: proof (induct K) wenzelm@18730: case empty berghofe@23751: from M0 show "M0 + {#} \ ?W" by simp wenzelm@18730: next wenzelm@18730: case (add K x) berghofe@23751: from add.prems have "(x, a) \ r" by simp berghofe@23751: with wf_hyp have "\M \ ?W. M + {#x#} \ ?W" by blast berghofe@23751: moreover from add have "M0 + K \ ?W" by simp berghofe@23751: ultimately have "(M0 + K) + {#x#} \ ?W" .. berghofe@23751: then show "M0 + (K + {#x#}) \ ?W" by (simp only: union_assoc) wenzelm@10249: qed berghofe@23751: then show "N \ ?W" by (simp only: N) wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: } note tedious_reasoning = this wenzelm@10249: berghofe@23751: assume wf: "wf r" wenzelm@10249: fix M berghofe@23751: show "M \ ?W" wenzelm@10249: proof (induct M) berghofe@23751: show "{#} \ ?W" wenzelm@10249: proof (rule accI) berghofe@23751: fix b assume "(b, {#}) \ ?R" berghofe@23751: with not_less_empty show "b \ ?W" by contradiction wenzelm@10249: qed wenzelm@10249: berghofe@23751: fix M a assume "M \ ?W" berghofe@23751: from wf have "\M \ ?W. M + {#a#} \ ?W" wenzelm@10249: proof induct wenzelm@10249: fix a berghofe@23751: assume r: "!!b. (b, a) \ r ==> (\M \ ?W. M + {#b#} \ ?W)" berghofe@23751: show "\M \ ?W. M + {#a#} \ ?W" wenzelm@10249: proof berghofe@23751: fix M assume "M \ ?W" berghofe@23751: then show "M + {#a#} \ ?W" wenzelm@23373: by (rule acc_induct) (rule tedious_reasoning [OF _ r]) wenzelm@10249: qed wenzelm@10249: qed berghofe@23751: from this and `M \ ?W` show "M + {#a#} \ ?W" .. wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: berghofe@23751: theorem wf_mult1: "wf r ==> wf (mult1 r)" nipkow@26178: by (rule acc_wfI) (rule all_accessible) wenzelm@10249: berghofe@23751: theorem wf_mult: "wf r ==> wf (mult r)" nipkow@26178: unfolding mult_def by (rule wf_trancl) (rule wf_mult1) wenzelm@10249: wenzelm@10249: wenzelm@10249: subsubsection {* Closure-free presentation *} wenzelm@10249: wenzelm@10249: (*Badly needed: a linear arithmetic procedure for multisets*) wenzelm@10249: wenzelm@10249: lemma diff_union_single_conv: "a :# J ==> I + J - {#a#} = I + (J - {#a#})" nipkow@26178: by (simp add: multiset_eq_conv_count_eq) wenzelm@10249: wenzelm@10249: text {* One direction. *} wenzelm@10249: wenzelm@10249: lemma mult_implies_one_step: berghofe@23751: "trans r ==> (M, N) \ mult r ==> nipkow@11464: \I J K. N = I + J \ M = I + K \ J \ {#} \ berghofe@23751: (\k \ set_of K. \j \ set_of J. (k, j) \ r)" nipkow@26178: apply (unfold mult_def mult1_def set_of_def) nipkow@26178: apply (erule converse_trancl_induct, clarify) nipkow@26178: apply (rule_tac x = M0 in exI, simp, clarify) nipkow@26178: apply (case_tac "a :# K") nipkow@26178: apply (rule_tac x = I in exI) nipkow@26178: apply (simp (no_asm)) nipkow@26178: apply (rule_tac x = "(K - {#a#}) + Ka" in exI) nipkow@26178: apply (simp (no_asm_simp) add: union_assoc [symmetric]) nipkow@26178: apply (drule_tac f = "\M. M - {#a#}" in arg_cong) nipkow@26178: apply (simp add: diff_union_single_conv) nipkow@26178: apply (simp (no_asm_use) add: trans_def) nipkow@26178: apply blast nipkow@26178: apply (subgoal_tac "a :# I") nipkow@26178: apply (rule_tac x = "I - {#a#}" in exI) nipkow@26178: apply (rule_tac x = "J + {#a#}" in exI) nipkow@26178: apply (rule_tac x = "K + Ka" in exI) nipkow@26178: apply (rule conjI) nipkow@26178: apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) nipkow@26178: apply (rule conjI) nipkow@26178: apply (drule_tac f = "\M. M - {#a#}" in arg_cong, simp) nipkow@26178: apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) nipkow@26178: apply (simp (no_asm_use) add: trans_def) nipkow@26178: apply blast nipkow@26178: apply (subgoal_tac "a :# (M0 + {#a#})") nipkow@26178: apply simp nipkow@26178: apply (simp (no_asm)) nipkow@26178: done wenzelm@10249: wenzelm@10249: lemma elem_imp_eq_diff_union: "a :# M ==> M = M - {#a#} + {#a#}" nipkow@26178: by (simp add: multiset_eq_conv_count_eq) wenzelm@10249: nipkow@11464: lemma size_eq_Suc_imp_eq_union: "size M = Suc n ==> \a N. M = N + {#a#}" nipkow@26178: apply (erule size_eq_Suc_imp_elem [THEN exE]) nipkow@26178: apply (drule elem_imp_eq_diff_union, auto) nipkow@26178: done wenzelm@10249: wenzelm@10249: lemma one_step_implies_mult_aux: berghofe@23751: "trans r ==> berghofe@23751: \I J K. (size J = n \ J \ {#} \ (\k \ set_of K. \j \ set_of J. (k, j) \ r)) berghofe@23751: --> (I + K, I + J) \ mult r" nipkow@26178: apply (induct_tac n, auto) nipkow@26178: apply (frule size_eq_Suc_imp_eq_union, clarify) nipkow@26178: apply (rename_tac "J'", simp) nipkow@26178: apply (erule notE, auto) nipkow@26178: apply (case_tac "J' = {#}") nipkow@26178: apply (simp add: mult_def) nipkow@26178: apply (rule r_into_trancl) nipkow@26178: apply (simp add: mult1_def set_of_def, blast) nipkow@26178: txt {* Now we know @{term "J' \ {#}"}. *} nipkow@26178: apply (cut_tac M = K and P = "\x. (x, a) \ r" in multiset_partition) nipkow@26178: apply (erule_tac P = "\k \ set_of K. ?P k" in rev_mp) nipkow@26178: apply (erule ssubst) nipkow@26178: apply (simp add: Ball_def, auto) nipkow@26178: apply (subgoal_tac nipkow@26178: "((I + {# x :# K. (x, a) \ r #}) + {# x :# K. (x, a) \ r #}, nipkow@26178: (I + {# x :# K. (x, a) \ r #}) + J') \ mult r") nipkow@26178: prefer 2 nipkow@26178: apply force nipkow@26178: apply (simp (no_asm_use) add: union_assoc [symmetric] mult_def) nipkow@26178: apply (erule trancl_trans) nipkow@26178: apply (rule r_into_trancl) nipkow@26178: apply (simp add: mult1_def set_of_def) nipkow@26178: apply (rule_tac x = a in exI) nipkow@26178: apply (rule_tac x = "I + J'" in exI) nipkow@26178: apply (simp add: union_ac) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma one_step_implies_mult: berghofe@23751: "trans r ==> J \ {#} ==> \k \ set_of K. \j \ set_of J. (k, j) \ r berghofe@23751: ==> (I + K, I + J) \ mult r" nipkow@26178: using one_step_implies_mult_aux by blast wenzelm@10249: wenzelm@10249: wenzelm@10249: subsubsection {* Partial-order properties *} wenzelm@10249: haftmann@26567: instantiation multiset :: (order) order haftmann@26567: begin wenzelm@10249: haftmann@28708: definition less_multiset_def [code del]: haftmann@28708: "M' < M \ (M', M) \ mult {(x', x). x' < x}" haftmann@26567: haftmann@28708: definition le_multiset_def [code del]: haftmann@28708: "M' <= M \ M' = M \ M' < (M::'a multiset)" wenzelm@10249: berghofe@23751: lemma trans_base_order: "trans {(x', x). x' < (x::'a::order)}" nipkow@26178: unfolding trans_def by (blast intro: order_less_trans) wenzelm@10249: wenzelm@10249: text {* wenzelm@10249: \medskip Irreflexivity. wenzelm@10249: *} wenzelm@10249: wenzelm@10249: lemma mult_irrefl_aux: nipkow@26178: "finite A ==> (\x \ A. \y \ A. x < (y::'a::order)) \ A = {}" nipkow@26178: by (induct rule: finite_induct) (auto intro: order_less_trans) wenzelm@10249: wenzelm@17161: lemma mult_less_not_refl: "\ M < (M::'a::order multiset)" nipkow@26178: apply (unfold less_multiset_def, auto) nipkow@26178: apply (drule trans_base_order [THEN mult_implies_one_step], auto) nipkow@26178: apply (drule finite_set_of [THEN mult_irrefl_aux [rule_format (no_asm)]]) nipkow@26178: apply (simp add: set_of_eq_empty_iff) nipkow@26178: done wenzelm@10249: wenzelm@10249: lemma mult_less_irrefl [elim!]: "M < (M::'a::order multiset) ==> R" nipkow@26178: using insert mult_less_not_refl by fast wenzelm@10249: wenzelm@10249: wenzelm@10249: text {* Transitivity. *} wenzelm@10249: wenzelm@10249: theorem mult_less_trans: "K < M ==> M < N ==> K < (N::'a::order multiset)" nipkow@26178: unfolding less_multiset_def mult_def by (blast intro: trancl_trans) wenzelm@10249: wenzelm@10249: text {* Asymmetry. *} wenzelm@10249: nipkow@11464: theorem mult_less_not_sym: "M < N ==> \ N < (M::'a::order multiset)" nipkow@26178: apply auto nipkow@26178: apply (rule mult_less_not_refl [THEN notE]) nipkow@26178: apply (erule mult_less_trans, assumption) nipkow@26178: done wenzelm@10249: wenzelm@10249: theorem mult_less_asym: nipkow@26178: "M < N ==> (\ P ==> N < (M::'a::order multiset)) ==> P" nipkow@26178: using mult_less_not_sym by blast wenzelm@10249: wenzelm@10249: theorem mult_le_refl [iff]: "M <= (M::'a::order multiset)" nipkow@26178: unfolding le_multiset_def by auto wenzelm@10249: wenzelm@10249: text {* Anti-symmetry. *} wenzelm@10249: wenzelm@10249: theorem mult_le_antisym: nipkow@26178: "M <= N ==> N <= M ==> M = (N::'a::order multiset)" nipkow@26178: unfolding le_multiset_def by (blast dest: mult_less_not_sym) wenzelm@10249: wenzelm@10249: text {* Transitivity. *} wenzelm@10249: wenzelm@10249: theorem mult_le_trans: nipkow@26178: "K <= M ==> M <= N ==> K <= (N::'a::order multiset)" nipkow@26178: unfolding le_multiset_def by (blast intro: mult_less_trans) wenzelm@10249: wenzelm@11655: theorem mult_less_le: "(M < N) = (M <= N \ M \ (N::'a::order multiset))" nipkow@26178: unfolding le_multiset_def by auto wenzelm@10249: haftmann@27682: instance proof haftmann@27682: qed (auto simp add: mult_less_le dest: mult_le_antisym elim: mult_le_trans) wenzelm@10277: haftmann@26567: end haftmann@26567: wenzelm@10249: wenzelm@10249: subsubsection {* Monotonicity of multiset union *} wenzelm@10249: wenzelm@17161: lemma mult1_union: nipkow@26178: "(B, D) \ mult1 r ==> trans r ==> (C + B, C + D) \ mult1 r" nipkow@26178: apply (unfold mult1_def) nipkow@26178: apply auto nipkow@26178: apply (rule_tac x = a in exI) nipkow@26178: apply (rule_tac x = "C + M0" in exI) nipkow@26178: apply (simp add: union_assoc) nipkow@26178: done wenzelm@10249: wenzelm@10249: lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)" nipkow@26178: apply (unfold less_multiset_def mult_def) nipkow@26178: apply (erule trancl_induct) nipkow@26178: apply (blast intro: mult1_union transI order_less_trans r_into_trancl) nipkow@26178: apply (blast intro: mult1_union transI order_less_trans r_into_trancl trancl_trans) nipkow@26178: done wenzelm@10249: wenzelm@10249: lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)" nipkow@26178: apply (subst union_commute [of B C]) nipkow@26178: apply (subst union_commute [of D C]) nipkow@26178: apply (erule union_less_mono2) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma union_less_mono: nipkow@26178: "A < C ==> B < D ==> A + B < C + (D::'a::order multiset)" nipkow@26178: by (blast intro!: union_less_mono1 union_less_mono2 mult_less_trans) wenzelm@10249: wenzelm@17161: lemma union_le_mono: nipkow@26178: "A <= C ==> B <= D ==> A + B <= C + (D::'a::order multiset)" nipkow@26178: unfolding le_multiset_def nipkow@26178: by (blast intro: union_less_mono union_less_mono1 union_less_mono2) wenzelm@10249: wenzelm@17161: lemma empty_leI [iff]: "{#} <= (M::'a::order multiset)" nipkow@26178: apply (unfold le_multiset_def less_multiset_def) nipkow@26178: apply (case_tac "M = {#}") nipkow@26178: prefer 2 nipkow@26178: apply (subgoal_tac "({#} + {#}, {#} + M) \ mult (Collect (split op <))") nipkow@26178: prefer 2 nipkow@26178: apply (rule one_step_implies_mult) nipkow@26178: apply (simp only: trans_def) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma union_upper1: "A <= A + (B::'a::order multiset)" paulson@15072: proof - wenzelm@17200: have "A + {#} <= A + B" by (blast intro: union_le_mono) wenzelm@18258: then show ?thesis by simp paulson@15072: qed paulson@15072: wenzelm@17161: lemma union_upper2: "B <= A + (B::'a::order multiset)" nipkow@26178: by (subst union_commute) (rule union_upper1) paulson@15072: nipkow@23611: instance multiset :: (order) pordered_ab_semigroup_add nipkow@26178: apply intro_classes nipkow@26178: apply (erule union_le_mono[OF mult_le_refl]) nipkow@26178: done wenzelm@26145: paulson@15072: wenzelm@17200: subsection {* Link with lists *} paulson@15072: nipkow@26016: primrec multiset_of :: "'a list \ 'a multiset" where wenzelm@26145: "multiset_of [] = {#}" | wenzelm@26145: "multiset_of (a # x) = multiset_of x + {# a #}" paulson@15072: paulson@15072: lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])" nipkow@26178: by (induct x) auto paulson@15072: paulson@15072: lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])" nipkow@26178: by (induct x) auto paulson@15072: paulson@15072: lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" nipkow@26178: by (induct x) auto kleing@15867: kleing@15867: lemma mem_set_multiset_eq: "x \ set xs = (x :# multiset_of xs)" nipkow@26178: by (induct xs) auto paulson@15072: wenzelm@18258: lemma multiset_of_append [simp]: nipkow@26178: "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" nipkow@26178: by (induct xs arbitrary: ys) (auto simp: union_ac) wenzelm@18730: paulson@15072: lemma surj_multiset_of: "surj multiset_of" nipkow@26178: apply (unfold surj_def) nipkow@26178: apply (rule allI) nipkow@26178: apply (rule_tac M = y in multiset_induct) nipkow@26178: apply auto nipkow@26178: apply (rule_tac x = "x # xa" in exI) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: nipkow@25162: lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}" nipkow@26178: by (induct x) auto paulson@15072: wenzelm@17200: lemma distinct_count_atmost_1: nipkow@26178: "distinct x = (! a. count (multiset_of x) a = (if a \ set x then 1 else 0))" nipkow@26178: apply (induct x, simp, rule iffI, simp_all) nipkow@26178: apply (rule conjI) nipkow@26178: apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) nipkow@26178: apply (erule_tac x = a in allE, simp, clarify) nipkow@26178: apply (erule_tac x = aa in allE, simp) nipkow@26178: done paulson@15072: wenzelm@17200: lemma multiset_of_eq_setD: kleing@15867: "multiset_of xs = multiset_of ys \ set xs = set ys" nipkow@26178: by (rule) (auto simp add:multiset_eq_conv_count_eq set_count_greater_0) kleing@15867: wenzelm@17200: lemma set_eq_iff_multiset_of_eq_distinct: wenzelm@26145: "distinct x \ distinct y \ wenzelm@26145: (set x = set y) = (multiset_of x = multiset_of y)" nipkow@26178: by (auto simp: multiset_eq_conv_count_eq distinct_count_atmost_1) paulson@15072: wenzelm@17200: lemma set_eq_iff_multiset_of_remdups_eq: paulson@15072: "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" nipkow@26178: apply (rule iffI) nipkow@26178: apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) nipkow@26178: apply (drule distinct_remdups [THEN distinct_remdups wenzelm@26145: [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]]) nipkow@26178: apply simp nipkow@26178: done wenzelm@10249: wenzelm@18258: lemma multiset_of_compl_union [simp]: nipkow@26178: "multiset_of [x\xs. P x] + multiset_of [x\xs. \P x] = multiset_of xs" nipkow@26178: by (induct xs) (auto simp: union_ac) paulson@15072: wenzelm@17200: lemma count_filter: nipkow@26178: "count (multiset_of xs) x = length [y \ xs. y = x]" nipkow@26178: by (induct xs) auto kleing@15867: bulwahn@26143: lemma nth_mem_multiset_of: "i < length ls \ (ls ! i) :# multiset_of ls" nipkow@26178: apply (induct ls arbitrary: i) nipkow@26178: apply simp nipkow@26178: apply (case_tac i) nipkow@26178: apply auto nipkow@26178: done bulwahn@26143: bulwahn@26143: lemma multiset_of_remove1: "multiset_of (remove1 a xs) = multiset_of xs - {#a#}" nipkow@26178: by (induct xs) (auto simp add: multiset_eq_conv_count_eq) bulwahn@26143: bulwahn@26143: lemma multiset_of_eq_length: nipkow@26178: assumes "multiset_of xs = multiset_of ys" nipkow@26178: shows "length xs = length ys" nipkow@26178: using assms bulwahn@26143: proof (induct arbitrary: ys rule: length_induct) bulwahn@26143: case (1 xs ys) bulwahn@26143: show ?case bulwahn@26143: proof (cases xs) wenzelm@26145: case Nil with "1.prems" show ?thesis by simp bulwahn@26143: next bulwahn@26143: case (Cons x xs') bulwahn@26143: note xCons = Cons bulwahn@26143: show ?thesis bulwahn@26143: proof (cases ys) bulwahn@26143: case Nil wenzelm@26145: with "1.prems" Cons show ?thesis by simp bulwahn@26143: next bulwahn@26143: case (Cons y ys') bulwahn@26143: have x_in_ys: "x = y \ x \ set ys'" bulwahn@26143: proof (cases "x = y") wenzelm@26145: case True then show ?thesis .. bulwahn@26143: next bulwahn@26143: case False wenzelm@26145: from "1.prems" [symmetric] xCons Cons have "x :# multiset_of ys' + {#y#}" by simp bulwahn@26143: with False show ?thesis by (simp add: mem_set_multiset_eq) bulwahn@26143: qed wenzelm@26145: from "1.hyps" have IH: "length xs' < length xs \ wenzelm@26145: (\x. multiset_of xs' = multiset_of x \ length xs' = length x)" by blast wenzelm@26145: from "1.prems" x_in_ys Cons xCons have "multiset_of xs' = multiset_of (remove1 x (y#ys'))" bulwahn@26143: apply - bulwahn@26143: apply (simp add: multiset_of_remove1, simp only: add_eq_conv_diff) bulwahn@26143: apply fastsimp bulwahn@26143: done wenzelm@26145: with IH xCons have IH': "length xs' = length (remove1 x (y#ys'))" by fastsimp wenzelm@26145: from x_in_ys have "x \ y \ length ys' > 0" by auto bulwahn@26143: with Cons xCons x_in_ys IH' show ?thesis by (auto simp add: length_remove1) bulwahn@26143: qed bulwahn@26143: qed bulwahn@26143: qed bulwahn@26143: wenzelm@26145: text {* wenzelm@26145: This lemma shows which properties suffice to show that a function wenzelm@26145: @{text "f"} with @{text "f xs = ys"} behaves like sort. wenzelm@26145: *} wenzelm@26145: lemma properties_for_sort: wenzelm@26145: "multiset_of ys = multiset_of xs \ sorted ys \ sort xs = ys" bulwahn@26143: proof (induct xs arbitrary: ys) wenzelm@26145: case Nil then show ?case by simp bulwahn@26143: next bulwahn@26143: case (Cons x xs) wenzelm@26145: then have "x \ set ys" wenzelm@26145: by (auto simp add: mem_set_multiset_eq intro!: ccontr) bulwahn@26143: with Cons.prems Cons.hyps [of "remove1 x ys"] show ?case bulwahn@26143: by (simp add: sorted_remove1 multiset_of_remove1 insort_remove1) bulwahn@26143: qed bulwahn@26143: kleing@15867: paulson@15072: subsection {* Pointwise ordering induced by count *} paulson@15072: haftmann@28708: definition mset_le :: "'a multiset \ 'a multiset \ bool" (infix "\#" 50) where haftmann@28708: [code del]: "A \# B \ (\a. count A a \ count B a)" wenzelm@26145: haftmann@28708: definition mset_less :: "'a multiset \ 'a multiset \ bool" (infix "<#" 50) where haftmann@28708: [code del]: "A <# B \ A \# B \ A \ B" kleing@25610: wenzelm@26145: notation mset_le (infix "\#" 50) wenzelm@26145: notation mset_less (infix "\#" 50) paulson@15072: nipkow@23611: lemma mset_le_refl[simp]: "A \# A" nipkow@26178: unfolding mset_le_def by auto paulson@15072: wenzelm@26145: lemma mset_le_trans: "A \# B \ B \# C \ A \# C" nipkow@26178: unfolding mset_le_def by (fast intro: order_trans) paulson@15072: wenzelm@26145: lemma mset_le_antisym: "A \# B \ B \# A \ A = B" nipkow@26178: apply (unfold mset_le_def) nipkow@26178: apply (rule multiset_eq_conv_count_eq [THEN iffD2]) nipkow@26178: apply (blast intro: order_antisym) nipkow@26178: done paulson@15072: wenzelm@26145: lemma mset_le_exists_conv: "(A \# B) = (\C. B = A + C)" nipkow@26178: apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI) nipkow@26178: apply (auto intro: multiset_eq_conv_count_eq [THEN iffD2]) nipkow@26178: done paulson@15072: nipkow@23611: lemma mset_le_mono_add_right_cancel[simp]: "(A + C \# B + C) = (A \# B)" nipkow@26178: unfolding mset_le_def by auto paulson@15072: nipkow@23611: lemma mset_le_mono_add_left_cancel[simp]: "(C + A \# C + B) = (A \# B)" nipkow@26178: unfolding mset_le_def by auto paulson@15072: nipkow@23611: lemma mset_le_mono_add: "\ A \# B; C \# D \ \ A + C \# B + D" nipkow@26178: apply (unfold mset_le_def) nipkow@26178: apply auto nipkow@26178: apply (erule_tac x = a in allE)+ nipkow@26178: apply auto nipkow@26178: done paulson@15072: nipkow@23611: lemma mset_le_add_left[simp]: "A \# A + B" nipkow@26178: unfolding mset_le_def by auto paulson@15072: nipkow@23611: lemma mset_le_add_right[simp]: "B \# A + B" nipkow@26178: unfolding mset_le_def by auto paulson@15072: bulwahn@26143: lemma mset_le_single: "a :# B \ {#a#} \# B" nipkow@26178: by (simp add: mset_le_def) bulwahn@26143: bulwahn@26143: lemma multiset_diff_union_assoc: "C \# B \ A + B - C = A + (B - C)" nipkow@26178: by (simp add: multiset_eq_conv_count_eq mset_le_def) bulwahn@26143: bulwahn@26143: lemma mset_le_multiset_union_diff_commute: nipkow@26178: assumes "B \# A" nipkow@26178: shows "A - B + C = A + C - B" bulwahn@26143: proof - wenzelm@26145: from mset_le_exists_conv [of "B" "A"] assms have "\D. A = B + D" .. wenzelm@26145: from this obtain D where "A = B + D" .. wenzelm@26145: then show ?thesis wenzelm@26145: apply simp wenzelm@26145: apply (subst union_commute) wenzelm@26145: apply (subst multiset_diff_union_assoc) wenzelm@26145: apply simp wenzelm@26145: apply (simp add: diff_cancel) wenzelm@26145: apply (subst union_assoc) wenzelm@26145: apply (subst union_commute[of "B" _]) wenzelm@26145: apply (subst multiset_diff_union_assoc) wenzelm@26145: apply simp wenzelm@26145: apply (simp add: diff_cancel) wenzelm@26145: done bulwahn@26143: qed bulwahn@26143: nipkow@23611: lemma multiset_of_remdups_le: "multiset_of (remdups xs) \# multiset_of xs" nipkow@26178: apply (induct xs) nipkow@26178: apply auto nipkow@26178: apply (rule mset_le_trans) nipkow@26178: apply auto nipkow@26178: done nipkow@23611: wenzelm@26145: lemma multiset_of_update: wenzelm@26145: "i < length ls \ multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}" bulwahn@26143: proof (induct ls arbitrary: i) wenzelm@26145: case Nil then show ?case by simp bulwahn@26143: next bulwahn@26143: case (Cons x xs) bulwahn@26143: show ?case wenzelm@26145: proof (cases i) wenzelm@26145: case 0 then show ?thesis by simp wenzelm@26145: next wenzelm@26145: case (Suc i') wenzelm@26145: with Cons show ?thesis wenzelm@26145: apply simp wenzelm@26145: apply (subst union_assoc) wenzelm@26145: apply (subst union_commute [where M = "{#v#}" and N = "{#x#}"]) wenzelm@26145: apply (subst union_assoc [symmetric]) wenzelm@26145: apply simp wenzelm@26145: apply (rule mset_le_multiset_union_diff_commute) wenzelm@26145: apply (simp add: mset_le_single nth_mem_multiset_of) wenzelm@26145: done bulwahn@26143: qed bulwahn@26143: qed bulwahn@26143: wenzelm@26145: lemma multiset_of_swap: wenzelm@26145: "i < length ls \ j < length ls \ wenzelm@26145: multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls" nipkow@26178: apply (case_tac "i = j") nipkow@26178: apply simp nipkow@26178: apply (simp add: multiset_of_update) nipkow@26178: apply (subst elem_imp_eq_diff_union[symmetric]) nipkow@26178: apply (simp add: nth_mem_multiset_of) nipkow@26178: apply simp nipkow@26178: done bulwahn@26143: wenzelm@26145: interpretation mset_order: order ["op \#" "op <#"] haftmann@27682: proof qed (auto intro: order.intro mset_le_refl mset_le_antisym haftmann@27682: mset_le_trans simp: mset_less_def) nipkow@23611: nipkow@23611: interpretation mset_order_cancel_semigroup: haftmann@28823: pordered_cancel_ab_semigroup_add ["op +" "op \#" "op <#"] haftmann@27682: proof qed (erule mset_le_mono_add [OF mset_le_refl]) nipkow@23611: nipkow@23611: interpretation mset_order_semigroup_cancel: haftmann@28823: pordered_ab_semigroup_add_imp_le ["op +" "op \#" "op <#"] haftmann@27682: proof qed simp paulson@15072: kleing@25610: wenzelm@26145: lemma mset_lessD: "A \# B \ x \# A \ x \# B" nipkow@26178: apply (clarsimp simp: mset_le_def mset_less_def) nipkow@26178: apply (erule_tac x=x in allE) nipkow@26178: apply auto nipkow@26178: done kleing@25610: wenzelm@26145: lemma mset_leD: "A \# B \ x \# A \ x \# B" nipkow@26178: apply (clarsimp simp: mset_le_def mset_less_def) nipkow@26178: apply (erule_tac x = x in allE) nipkow@26178: apply auto nipkow@26178: done kleing@25610: wenzelm@26145: lemma mset_less_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" nipkow@26178: apply (rule conjI) nipkow@26178: apply (simp add: mset_lessD) nipkow@26178: apply (clarsimp simp: mset_le_def mset_less_def) nipkow@26178: apply safe nipkow@26178: apply (erule_tac x = a in allE) nipkow@26178: apply (auto split: split_if_asm) nipkow@26178: done kleing@25610: wenzelm@26145: lemma mset_le_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" nipkow@26178: apply (rule conjI) nipkow@26178: apply (simp add: mset_leD) nipkow@26178: apply (force simp: mset_le_def mset_less_def split: split_if_asm) nipkow@26178: done kleing@25610: kleing@25610: lemma mset_less_of_empty[simp]: "A \# {#} = False" nipkow@26178: by (induct A) (auto simp: mset_le_def mset_less_def) kleing@25610: kleing@25610: lemma multi_psub_of_add_self[simp]: "A \# A + {#x#}" nipkow@26178: by (auto simp: mset_le_def mset_less_def) kleing@25610: kleing@25610: lemma multi_psub_self[simp]: "A \# A = False" nipkow@26178: by (auto simp: mset_le_def mset_less_def) kleing@25610: kleing@25610: lemma mset_less_add_bothsides: kleing@25610: "T + {#x#} \# S + {#x#} \ T \# S" nipkow@26178: by (auto simp: mset_le_def mset_less_def) kleing@25610: kleing@25610: lemma mset_less_empty_nonempty: "({#} \# S) = (S \ {#})" nipkow@26178: by (auto simp: mset_le_def mset_less_def) kleing@25610: kleing@25610: lemma mset_less_size: "A \# B \ size A < size B" kleing@25610: proof (induct A arbitrary: B) kleing@25610: case (empty M) wenzelm@26145: then have "M \ {#}" by (simp add: mset_less_empty_nonempty) kleing@25610: then obtain M' x where "M = M' + {#x#}" kleing@25610: by (blast dest: multi_nonempty_split) wenzelm@26145: then show ?case by simp kleing@25610: next kleing@25610: case (add S x T) kleing@25610: have IH: "\B. S \# B \ size S < size B" by fact kleing@25610: have SxsubT: "S + {#x#} \# T" by fact wenzelm@26145: then have "x \# T" and "S \# T" by (auto dest: mset_less_insertD) kleing@25610: then obtain T' where T: "T = T' + {#x#}" kleing@25610: by (blast dest: multi_member_split) wenzelm@26145: then have "S \# T'" using SxsubT kleing@25610: by (blast intro: mset_less_add_bothsides) wenzelm@26145: then have "size S < size T'" using IH by simp wenzelm@26145: then show ?case using T by simp kleing@25610: qed kleing@25610: kleing@25610: lemmas mset_less_trans = mset_order.less_eq_less.less_trans kleing@25610: kleing@25610: lemma mset_less_diff_self: "c \# B \ B - {#c#} \# B" nipkow@26178: by (auto simp: mset_le_def mset_less_def multi_drop_mem_not_eq) kleing@25610: wenzelm@26145: kleing@25610: subsection {* Strong induction and subset induction for multisets *} kleing@25610: nipkow@26016: text {* Well-foundedness of proper subset operator: *} kleing@25610: wenzelm@26145: text {* proper multiset subset *} kleing@25610: definition wenzelm@26145: mset_less_rel :: "('a multiset * 'a multiset) set" where wenzelm@26145: "mset_less_rel = {(A,B). A \# B}" kleing@25610: kleing@25610: lemma multiset_add_sub_el_shuffle: wenzelm@26145: assumes "c \# B" and "b \ c" kleing@25610: shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}" kleing@25610: proof - wenzelm@26145: from `c \# B` obtain A where B: "B = A + {#c#}" kleing@25610: by (blast dest: multi_member_split) kleing@25610: have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp wenzelm@26145: then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" kleing@25610: by (simp add: union_ac) wenzelm@26145: then show ?thesis using B by simp kleing@25610: qed kleing@25610: kleing@25610: lemma wf_mset_less_rel: "wf mset_less_rel" nipkow@26178: apply (unfold mset_less_rel_def) nipkow@26178: apply (rule wf_measure [THEN wf_subset, where f1=size]) nipkow@26178: apply (clarsimp simp: measure_def inv_image_def mset_less_size) nipkow@26178: done kleing@25610: nipkow@26016: text {* The induction rules: *} kleing@25610: kleing@25610: lemma full_multiset_induct [case_names less]: nipkow@26178: assumes ih: "\B. \A. A \# B \ P A \ P B" nipkow@26178: shows "P B" nipkow@26178: apply (rule wf_mset_less_rel [THEN wf_induct]) nipkow@26178: apply (rule ih, auto simp: mset_less_rel_def) nipkow@26178: done kleing@25610: kleing@25610: lemma multi_subset_induct [consumes 2, case_names empty add]: nipkow@26178: assumes "F \# A" nipkow@26178: and empty: "P {#}" nipkow@26178: and insert: "\a F. a \# A \ P F \ P (F + {#a#})" nipkow@26178: shows "P F" kleing@25610: proof - kleing@25610: from `F \# A` kleing@25610: show ?thesis kleing@25610: proof (induct F) kleing@25610: show "P {#}" by fact kleing@25610: next kleing@25610: fix x F kleing@25610: assume P: "F \# A \ P F" and i: "F + {#x#} \# A" kleing@25610: show "P (F + {#x#})" kleing@25610: proof (rule insert) kleing@25610: from i show "x \# A" by (auto dest: mset_le_insertD) wenzelm@26145: from i have "F \# A" by (auto dest: mset_le_insertD) kleing@25610: with P show "P F" . kleing@25610: qed kleing@25610: qed kleing@25610: qed kleing@25610: nipkow@26016: text{* A consequence: Extensionality. *} kleing@25610: wenzelm@26145: lemma multi_count_eq: "(\x. count A x = count B x) = (A = B)" nipkow@26178: apply (rule iffI) nipkow@26178: prefer 2 nipkow@26178: apply clarsimp nipkow@26178: apply (induct A arbitrary: B rule: full_multiset_induct) nipkow@26178: apply (rename_tac C) nipkow@26178: apply (case_tac B rule: multiset_cases) nipkow@26178: apply (simp add: empty_multiset_count) nipkow@26178: apply simp nipkow@26178: apply (case_tac "x \# C") nipkow@26178: apply (force dest: multi_member_split) nipkow@26178: apply (erule_tac x = x in allE) nipkow@26178: apply simp nipkow@26178: done kleing@25610: kleing@25610: lemmas multi_count_ext = multi_count_eq [THEN iffD1, rule_format] kleing@25610: wenzelm@26145: kleing@25610: subsection {* The fold combinator *} kleing@25610: wenzelm@26145: text {* wenzelm@26145: The intended behaviour is wenzelm@26145: @{text "fold_mset f z {#x\<^isub>1, ..., x\<^isub>n#} = f x\<^isub>1 (\ (f x\<^isub>n z)\)"} wenzelm@26145: if @{text f} is associative-commutative. kleing@25610: *} kleing@25610: wenzelm@26145: text {* wenzelm@26145: The graph of @{text "fold_mset"}, @{text "z"}: the start element, wenzelm@26145: @{text "f"}: folding function, @{text "A"}: the multiset, @{text wenzelm@26145: "y"}: the result. wenzelm@26145: *} kleing@25610: inductive kleing@25759: fold_msetG :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b \ bool" kleing@25610: for f :: "'a \ 'b \ 'b" kleing@25610: and z :: 'b kleing@25610: where kleing@25759: emptyI [intro]: "fold_msetG f z {#} z" kleing@25759: | insertI [intro]: "fold_msetG f z A y \ fold_msetG f z (A + {#x#}) (f x y)" kleing@25610: kleing@25759: inductive_cases empty_fold_msetGE [elim!]: "fold_msetG f z {#} x" kleing@25759: inductive_cases insert_fold_msetGE: "fold_msetG f z (A + {#}) y" kleing@25610: kleing@25610: definition wenzelm@26145: fold_mset :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b" where wenzelm@26145: "fold_mset f z A = (THE x. fold_msetG f z A x)" kleing@25610: kleing@25759: lemma Diff1_fold_msetG: wenzelm@26145: "fold_msetG f z (A - {#x#}) y \ x \# A \ fold_msetG f z A (f x y)" nipkow@26178: apply (frule_tac x = x in fold_msetG.insertI) nipkow@26178: apply auto nipkow@26178: done kleing@25610: kleing@25759: lemma fold_msetG_nonempty: "\x. fold_msetG f z A x" nipkow@26178: apply (induct A) nipkow@26178: apply blast nipkow@26178: apply clarsimp nipkow@26178: apply (drule_tac x = x in fold_msetG.insertI) nipkow@26178: apply auto nipkow@26178: done kleing@25610: kleing@25759: lemma fold_mset_empty[simp]: "fold_mset f z {#} = z" nipkow@26178: unfolding fold_mset_def by blast kleing@25610: kleing@25610: locale left_commutative = nipkow@26178: fixes f :: "'a => 'b => 'b" nipkow@26178: assumes left_commute: "f x (f y z) = f y (f x z)" wenzelm@26145: begin kleing@25610: wenzelm@26145: lemma fold_msetG_determ: wenzelm@26145: "fold_msetG f z A x \ fold_msetG f z A y \ y = x" kleing@25610: proof (induct arbitrary: x y z rule: full_multiset_induct) kleing@25610: case (less M x\<^isub>1 x\<^isub>2 Z) kleing@25610: have IH: "\A. A \# M \ kleing@25759: (\x x' x''. fold_msetG f x'' A x \ fold_msetG f x'' A x' kleing@25610: \ x' = x)" by fact kleing@25759: have Mfoldx\<^isub>1: "fold_msetG f Z M x\<^isub>1" and Mfoldx\<^isub>2: "fold_msetG f Z M x\<^isub>2" by fact+ kleing@25610: show ?case kleing@25759: proof (rule fold_msetG.cases [OF Mfoldx\<^isub>1]) kleing@25610: assume "M = {#}" and "x\<^isub>1 = Z" wenzelm@26145: then show ?case using Mfoldx\<^isub>2 by auto kleing@25610: next kleing@25610: fix B b u kleing@25759: assume "M = B + {#b#}" and "x\<^isub>1 = f b u" and Bu: "fold_msetG f Z B u" wenzelm@26145: then have MBb: "M = B + {#b#}" and x\<^isub>1: "x\<^isub>1 = f b u" by auto kleing@25610: show ?case kleing@25759: proof (rule fold_msetG.cases [OF Mfoldx\<^isub>2]) kleing@25610: assume "M = {#}" "x\<^isub>2 = Z" wenzelm@26145: then show ?case using Mfoldx\<^isub>1 by auto kleing@25610: next kleing@25610: fix C c v kleing@25759: assume "M = C + {#c#}" and "x\<^isub>2 = f c v" and Cv: "fold_msetG f Z C v" wenzelm@26145: then have MCc: "M = C + {#c#}" and x\<^isub>2: "x\<^isub>2 = f c v" by auto wenzelm@26145: then have CsubM: "C \# M" by simp kleing@25610: from MBb have BsubM: "B \# M" by simp kleing@25610: show ?case kleing@25610: proof cases kleing@25610: assume "b=c" kleing@25610: then moreover have "B = C" using MBb MCc by auto kleing@25610: ultimately show ?thesis using Bu Cv x\<^isub>1 x\<^isub>2 CsubM IH by auto kleing@25610: next kleing@25610: assume diff: "b \ c" kleing@25610: let ?D = "B - {#c#}" kleing@25610: have cinB: "c \# B" and binC: "b \# C" using MBb MCc diff kleing@25610: by (auto intro: insert_noteq_member dest: sym) kleing@25610: have "B - {#c#} \# B" using cinB by (rule mset_less_diff_self) wenzelm@26145: then have DsubM: "?D \# M" using BsubM by (blast intro: mset_less_trans) kleing@25610: from MBb MCc have "B + {#b#} = C + {#c#}" by blast wenzelm@26145: then have [simp]: "B + {#b#} - {#c#} = C" kleing@25610: using MBb MCc binC cinB by auto kleing@25610: have B: "B = ?D + {#c#}" and C: "C = ?D + {#b#}" kleing@25610: using MBb MCc diff binC cinB kleing@25610: by (auto simp: multiset_add_sub_el_shuffle) kleing@25759: then obtain d where Dfoldd: "fold_msetG f Z ?D d" kleing@25759: using fold_msetG_nonempty by iprover wenzelm@26145: then have "fold_msetG f Z B (f c d)" using cinB kleing@25759: by (rule Diff1_fold_msetG) wenzelm@26145: then have "f c d = u" using IH BsubM Bu by blast kleing@25610: moreover kleing@25759: have "fold_msetG f Z C (f b d)" using binC cinB diff Dfoldd kleing@25610: by (auto simp: multiset_add_sub_el_shuffle kleing@25759: dest: fold_msetG.insertI [where x=b]) wenzelm@26145: then have "f b d = v" using IH CsubM Cv by blast kleing@25610: ultimately show ?thesis using x\<^isub>1 x\<^isub>2 kleing@25610: by (auto simp: left_commute) kleing@25610: qed kleing@25610: qed kleing@25610: qed kleing@25610: qed kleing@25610: wenzelm@26145: lemma fold_mset_insert_aux: wenzelm@26145: "(fold_msetG f z (A + {#x#}) v) = kleing@25759: (\y. fold_msetG f z A y \ v = f x y)" nipkow@26178: apply (rule iffI) nipkow@26178: prefer 2 nipkow@26178: apply blast nipkow@26178: apply (rule_tac A=A and f=f in fold_msetG_nonempty [THEN exE, standard]) nipkow@26178: apply (blast intro: fold_msetG_determ) nipkow@26178: done kleing@25610: wenzelm@26145: lemma fold_mset_equality: "fold_msetG f z A y \ fold_mset f z A = y" nipkow@26178: unfolding fold_mset_def by (blast intro: fold_msetG_determ) kleing@25610: wenzelm@26145: lemma fold_mset_insert: nipkow@26178: "fold_mset f z (A + {#x#}) = f x (fold_mset f z A)" nipkow@26178: apply (simp add: fold_mset_def fold_mset_insert_aux union_commute) nipkow@26178: apply (rule the_equality) nipkow@26178: apply (auto cong add: conj_cong wenzelm@26145: simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) nipkow@26178: done kleing@25759: wenzelm@26145: lemma fold_mset_insert_idem: nipkow@26178: "fold_mset f z (A + {#a#}) = f a (fold_mset f z A)" nipkow@26178: apply (simp add: fold_mset_def fold_mset_insert_aux) nipkow@26178: apply (rule the_equality) nipkow@26178: apply (auto cong add: conj_cong wenzelm@26145: simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) nipkow@26178: done kleing@25610: wenzelm@26145: lemma fold_mset_commute: "f x (fold_mset f z A) = fold_mset f (f x z) A" nipkow@26178: by (induct A) (auto simp: fold_mset_insert left_commute [of x]) nipkow@26178: wenzelm@26145: lemma fold_mset_single [simp]: "fold_mset f z {#x#} = f x z" nipkow@26178: using fold_mset_insert [of z "{#}"] by simp kleing@25610: wenzelm@26145: lemma fold_mset_union [simp]: wenzelm@26145: "fold_mset f z (A+B) = fold_mset f (fold_mset f z A) B" kleing@25759: proof (induct A) wenzelm@26145: case empty then show ?case by simp kleing@25759: next wenzelm@26145: case (add A x) wenzelm@26145: have "A + {#x#} + B = (A+B) + {#x#}" by(simp add:union_ac) wenzelm@26145: then have "fold_mset f z (A + {#x#} + B) = f x (fold_mset f z (A + B))" wenzelm@26145: by (simp add: fold_mset_insert) wenzelm@26145: also have "\ = fold_mset f (fold_mset f z (A + {#x#})) B" wenzelm@26145: by (simp add: fold_mset_commute[of x,symmetric] add fold_mset_insert) wenzelm@26145: finally show ?case . kleing@25759: qed kleing@25759: wenzelm@26145: lemma fold_mset_fusion: ballarin@27611: assumes "left_commutative g" ballarin@27611: shows "(\x y. h (g x y) = f x (h y)) \ h (fold_mset g w A) = fold_mset f (h w) A" (is "PROP ?P") ballarin@27611: proof - ballarin@27611: interpret left_commutative [g] by fact ballarin@27611: show "PROP ?P" by (induct A) auto ballarin@27611: qed kleing@25610: wenzelm@26145: lemma fold_mset_rec: wenzelm@26145: assumes "a \# A" kleing@25759: shows "fold_mset f z A = f a (fold_mset f z (A - {#a#}))" kleing@25610: proof - wenzelm@26145: from assms obtain A' where "A = A' + {#a#}" wenzelm@26145: by (blast dest: multi_member_split) wenzelm@26145: then show ?thesis by simp kleing@25610: qed kleing@25610: wenzelm@26145: end wenzelm@26145: wenzelm@26145: text {* wenzelm@26145: A note on code generation: When defining some function containing a wenzelm@26145: subterm @{term"fold_mset F"}, code generation is not automatic. When wenzelm@26145: interpreting locale @{text left_commutative} with @{text F}, the wenzelm@26145: would be code thms for @{const fold_mset} become thms like wenzelm@26145: @{term"fold_mset F z {#} = z"} where @{text F} is not a pattern but wenzelm@26145: contains defined symbols, i.e.\ is not a code thm. Hence a separate wenzelm@26145: constant with its own code thms needs to be introduced for @{text wenzelm@26145: F}. See the image operator below. wenzelm@26145: *} wenzelm@26145: nipkow@26016: nipkow@26016: subsection {* Image *} nipkow@26016: haftmann@28708: definition [code del]: haftmann@28708: "image_mset f = fold_mset (op + o single o f) {#}" nipkow@26016: wenzelm@26145: interpretation image_left_comm: left_commutative ["op + o single o f"] haftmann@28823: proof qed (simp add:union_ac) nipkow@26016: haftmann@28708: lemma image_mset_empty [simp]: "image_mset f {#} = {#}" nipkow@26178: by (simp add: image_mset_def) nipkow@26016: haftmann@28708: lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}" nipkow@26178: by (simp add: image_mset_def) nipkow@26016: nipkow@26016: lemma image_mset_insert: nipkow@26016: "image_mset f (M + {#a#}) = image_mset f M + {#f a#}" nipkow@26178: by (simp add: image_mset_def add_ac) nipkow@26016: haftmann@28708: lemma image_mset_union [simp]: nipkow@26016: "image_mset f (M+N) = image_mset f M + image_mset f N" nipkow@26178: apply (induct N) nipkow@26178: apply simp nipkow@26178: apply (simp add: union_assoc [symmetric] image_mset_insert) nipkow@26178: done nipkow@26016: wenzelm@26145: lemma size_image_mset [simp]: "size (image_mset f M) = size M" nipkow@26178: by (induct M) simp_all nipkow@26016: wenzelm@26145: lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" nipkow@26178: by (cases M) auto nipkow@26016: wenzelm@26145: syntax wenzelm@26145: comprehension1_mset :: "'a \ 'b \ 'b multiset \ 'a multiset" wenzelm@26145: ("({#_/. _ :# _#})") wenzelm@26145: translations wenzelm@26145: "{#e. x:#M#}" == "CONST image_mset (%x. e) M" nipkow@26016: wenzelm@26145: syntax wenzelm@26145: comprehension2_mset :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" wenzelm@26145: ("({#_/ | _ :# _./ _#})") nipkow@26016: translations nipkow@26033: "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}" nipkow@26016: wenzelm@26145: text {* wenzelm@26145: This allows to write not just filters like @{term "{#x:#M. x# XS \ x \# {# y #} + XS" krauss@29125: and multi_member_this: "x \# {# x #} + XS" krauss@29125: and multi_member_last: "x \# {# x #}" krauss@29125: by auto krauss@29125: krauss@29125: definition "ms_strict = mult pair_less" krauss@29125: definition "ms_weak = ms_strict \ Id" krauss@29125: krauss@29125: lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)" krauss@29125: unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def krauss@29125: by (auto intro: wf_mult1 wf_trancl simp: mult_def) krauss@29125: krauss@29125: lemma smsI: krauss@29125: "(set_of A, set_of B) \ max_strict \ (Z + A, Z + B) \ ms_strict" krauss@29125: unfolding ms_strict_def krauss@29125: by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases) krauss@29125: krauss@29125: lemma wmsI: krauss@29125: "(set_of A, set_of B) \ max_strict \ A = {#} \ B = {#} krauss@29125: \ (Z + A, Z + B) \ ms_weak" krauss@29125: unfolding ms_weak_def ms_strict_def krauss@29125: by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult) krauss@29125: krauss@29125: inductive pw_leq krauss@29125: where krauss@29125: pw_leq_empty: "pw_leq {#} {#}" krauss@29125: | pw_leq_step: "\(x,y) \ pair_leq; pw_leq X Y \ \ pw_leq ({#x#} + X) ({#y#} + Y)" krauss@29125: krauss@29125: lemma pw_leq_lstep: krauss@29125: "(x, y) \ pair_leq \ pw_leq {#x#} {#y#}" krauss@29125: by (drule pw_leq_step) (rule pw_leq_empty, simp) krauss@29125: krauss@29125: lemma pw_leq_split: krauss@29125: assumes "pw_leq X Y" krauss@29125: shows "\A B Z. X = A + Z \ Y = B + Z \ ((set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#}))" krauss@29125: using assms krauss@29125: proof (induct) krauss@29125: case pw_leq_empty thus ?case by auto krauss@29125: next krauss@29125: case (pw_leq_step x y X Y) krauss@29125: then obtain A B Z where krauss@29125: [simp]: "X = A + Z" "Y = B + Z" krauss@29125: and 1[simp]: "(set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#})" krauss@29125: by auto krauss@29125: from pw_leq_step have "x = y \ (x, y) \ pair_less" krauss@29125: unfolding pair_leq_def by auto krauss@29125: thus ?case krauss@29125: proof krauss@29125: assume [simp]: "x = y" krauss@29125: have krauss@29125: "{#x#} + X = A + ({#y#}+Z) krauss@29125: \ {#y#} + Y = B + ({#y#}+Z) krauss@29125: \ ((set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#}))" krauss@29125: by (auto simp: add_ac) krauss@29125: thus ?case by (intro exI) krauss@29125: next krauss@29125: assume A: "(x, y) \ pair_less" krauss@29125: let ?A' = "{#x#} + A" and ?B' = "{#y#} + B" krauss@29125: have "{#x#} + X = ?A' + Z" krauss@29125: "{#y#} + Y = ?B' + Z" krauss@29125: by (auto simp add: add_ac) krauss@29125: moreover have krauss@29125: "(set_of ?A', set_of ?B') \ max_strict" krauss@29125: using 1 A unfolding max_strict_def krauss@29125: by (auto elim!: max_ext.cases) krauss@29125: ultimately show ?thesis by blast krauss@29125: qed krauss@29125: qed krauss@29125: krauss@29125: lemma krauss@29125: assumes pwleq: "pw_leq Z Z'" krauss@29125: shows ms_strictI: "(set_of A, set_of B) \ max_strict \ (Z + A, Z' + B) \ ms_strict" krauss@29125: and ms_weakI1: "(set_of A, set_of B) \ max_strict \ (Z + A, Z' + B) \ ms_weak" krauss@29125: and ms_weakI2: "(Z + {#}, Z' + {#}) \ ms_weak" krauss@29125: proof - krauss@29125: from pw_leq_split[OF pwleq] krauss@29125: obtain A' B' Z'' krauss@29125: where [simp]: "Z = A' + Z''" "Z' = B' + Z''" krauss@29125: and mx_or_empty: "(set_of A', set_of B') \ max_strict \ (A' = {#} \ B' = {#})" krauss@29125: by blast krauss@29125: { krauss@29125: assume max: "(set_of A, set_of B) \ max_strict" krauss@29125: from mx_or_empty krauss@29125: have "(Z'' + (A + A'), Z'' + (B + B')) \ ms_strict" krauss@29125: proof krauss@29125: assume max': "(set_of A', set_of B') \ max_strict" krauss@29125: with max have "(set_of (A + A'), set_of (B + B')) \ max_strict" krauss@29125: by (auto simp: max_strict_def intro: max_ext_additive) krauss@29125: thus ?thesis by (rule smsI) krauss@29125: next krauss@29125: assume [simp]: "A' = {#} \ B' = {#}" krauss@29125: show ?thesis by (rule smsI) (auto intro: max) krauss@29125: qed krauss@29125: thus "(Z + A, Z' + B) \ ms_strict" by (simp add:add_ac) krauss@29125: thus "(Z + A, Z' + B) \ ms_weak" by (simp add: ms_weak_def) krauss@29125: } krauss@29125: from mx_or_empty krauss@29125: have "(Z'' + A', Z'' + B') \ ms_weak" by (rule wmsI) krauss@29125: thus "(Z + {#}, Z' + {#}) \ ms_weak" by (simp add:add_ac) krauss@29125: qed krauss@29125: krauss@29125: lemma empty_idemp: "{#} + x = x" "x + {#} = x" krauss@29125: and nonempty_plus: "{# x #} + rs \ {#}" krauss@29125: and nonempty_single: "{# x #} \ {#}" krauss@29125: by auto krauss@29125: krauss@29125: setup {* krauss@29125: let krauss@29125: fun msetT T = Type ("Multiset.multiset", [T]); krauss@29125: krauss@29125: fun mk_mset T [] = Const (@{const_name Mempty}, msetT T) krauss@29125: | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x krauss@29125: | mk_mset T (x :: xs) = krauss@29125: Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $ krauss@29125: mk_mset T [x] $ mk_mset T xs krauss@29125: krauss@29125: fun mset_member_tac m i = krauss@29125: (if m <= 0 then krauss@29125: rtac @{thm multi_member_this} i ORELSE rtac @{thm multi_member_last} i krauss@29125: else krauss@29125: rtac @{thm multi_member_skip} i THEN mset_member_tac (m - 1) i) krauss@29125: krauss@29125: val mset_nonempty_tac = krauss@29125: rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single} krauss@29125: krauss@29125: val regroup_munion_conv = krauss@29125: FundefLib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus} krauss@29125: (map (fn t => t RS eq_reflection) (@{thms union_ac} @ @{thms empty_idemp})) krauss@29125: krauss@29125: fun unfold_pwleq_tac i = krauss@29125: (rtac @{thm pw_leq_step} i THEN (fn st => unfold_pwleq_tac (i + 1) st)) krauss@29125: ORELSE (rtac @{thm pw_leq_lstep} i) krauss@29125: ORELSE (rtac @{thm pw_leq_empty} i) krauss@29125: krauss@29125: val set_of_simps = [@{thm set_of_empty}, @{thm set_of_single}, @{thm set_of_union}, krauss@29125: @{thm Un_insert_left}, @{thm Un_empty_left}] krauss@29125: in krauss@29125: ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset krauss@29125: { krauss@29125: msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, krauss@29125: mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, krauss@29125: mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_of_simps, krauss@29125: smsI'=@{thm ms_strictI}, wmsI2''=@{thm ms_weakI2}, wmsI1=@{thm ms_weakI1}, krauss@29125: reduction_pair=@{thm ms_reduction_pair} krauss@29125: }) wenzelm@10249: end krauss@29125: *} krauss@29125: krauss@29125: end