wenzelm@10249: (* Title: HOL/Library/Multiset.thy paulson@15072: Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker blanchet@55129: Author: Andrei Popescu, TU Muenchen blanchet@59813: Author: Jasmin Blanchette, Inria, LORIA, MPII blanchet@59813: Author: Dmitriy Traytel, TU Muenchen blanchet@59813: Author: Mathias Fleury, MPII wenzelm@10249: *) wenzelm@10249: wenzelm@60500: section \(Finite) multisets\ wenzelm@10249: nipkow@15131: theory Multiset haftmann@51599: imports Main nipkow@15131: begin wenzelm@10249: wenzelm@60500: subsection \The type of multisets\ wenzelm@10249: wenzelm@60606: definition "multiset = {f :: 'a \ nat. finite {x. f x > 0}}" wenzelm@60606: wenzelm@60606: typedef 'a multiset = "multiset :: ('a \ nat) set" haftmann@34943: morphisms count Abs_multiset wenzelm@45694: unfolding multiset_def wenzelm@10249: proof wenzelm@45694: show "(\x. 0::nat) \ {f. finite {x. f x > 0}}" by simp wenzelm@10249: qed wenzelm@10249: bulwahn@47429: setup_lifting type_definition_multiset wenzelm@19086: wenzelm@60606: 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@60606: lemma multiset_eq_iff: "M = N \ (\a. count M a = count N a)" nipkow@39302: by (simp only: count_inject [symmetric] fun_eq_iff) haftmann@34943: wenzelm@60606: lemma multiset_eqI: "(\x. count A x = count B x) \ A = B" nipkow@39302: using multiset_eq_iff by auto haftmann@34943: wenzelm@60606: text \Preservation of the representing set @{term multiset}.\ wenzelm@60606: wenzelm@60606: lemma const0_in_multiset: "(\a. 0) \ multiset" haftmann@34943: by (simp add: multiset_def) haftmann@34943: wenzelm@60606: lemma only1_in_multiset: "(\b. if b = a then n else 0) \ multiset" haftmann@34943: by (simp add: multiset_def) haftmann@34943: wenzelm@60606: lemma union_preserves_multiset: "M \ multiset \ N \ multiset \ (\a. M a + N a) \ multiset" haftmann@34943: by (simp add: multiset_def) haftmann@34943: haftmann@34943: lemma diff_preserves_multiset: haftmann@34943: assumes "M \ multiset" haftmann@34943: shows "(\a. M a - N a) \ multiset" haftmann@34943: proof - haftmann@34943: have "{x. N x < M x} \ {x. 0 < M x}" haftmann@34943: by auto haftmann@34943: with assms show ?thesis haftmann@34943: by (auto simp add: multiset_def intro: finite_subset) haftmann@34943: qed haftmann@34943: haftmann@41069: lemma filter_preserves_multiset: haftmann@34943: assumes "M \ multiset" haftmann@34943: shows "(\x. if P x then M x else 0) \ multiset" haftmann@34943: proof - haftmann@34943: have "{x. (P x \ 0 < M x) \ P x} \ {x. 0 < M x}" haftmann@34943: by auto haftmann@34943: with assms show ?thesis haftmann@34943: by (auto simp add: multiset_def intro: finite_subset) haftmann@34943: qed haftmann@34943: haftmann@34943: lemmas in_multiset = const0_in_multiset only1_in_multiset haftmann@41069: union_preserves_multiset diff_preserves_multiset filter_preserves_multiset haftmann@34943: haftmann@34943: wenzelm@60500: subsection \Representing multisets\ wenzelm@60500: wenzelm@60500: text \Multiset enumeration\ haftmann@34943: huffman@48008: instantiation multiset :: (type) cancel_comm_monoid_add haftmann@25571: begin haftmann@25571: bulwahn@47429: lift_definition zero_multiset :: "'a multiset" is "\a. 0" bulwahn@47429: by (rule const0_in_multiset) haftmann@25571: haftmann@34943: abbreviation Mempty :: "'a multiset" ("{#}") where haftmann@34943: "Mempty \ 0" haftmann@25571: wenzelm@60606: lift_definition plus_multiset :: "'a multiset \ 'a multiset \ 'a multiset" is "\M N. (\a. M a + N a)" bulwahn@47429: by (rule union_preserves_multiset) haftmann@25571: wenzelm@60606: lift_definition minus_multiset :: "'a multiset \ 'a multiset \ 'a multiset" is "\ M N. \a. M a - N a" haftmann@59815: by (rule diff_preserves_multiset) haftmann@59815: huffman@48008: instance wenzelm@60678: by (standard; transfer; simp add: fun_eq_iff) haftmann@25571: haftmann@25571: end wenzelm@10249: wenzelm@60606: lift_definition single :: "'a \ 'a multiset" is "\a b. if b = a then 1 else 0" bulwahn@47429: by (rule only1_in_multiset) kleing@15869: wenzelm@26145: syntax wenzelm@60606: "_multiset" :: "args \ 'a multiset" ("{#(_)#}") nipkow@25507: translations nipkow@25507: "{#x, xs#}" == "{#x#} + {#xs#}" nipkow@25507: "{#x#}" == "CONST single x" nipkow@25507: haftmann@34943: lemma count_empty [simp]: "count {#} a = 0" bulwahn@47429: by (simp add: zero_multiset.rep_eq) wenzelm@10249: haftmann@34943: lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)" bulwahn@47429: by (simp add: single.rep_eq) nipkow@29901: wenzelm@10249: wenzelm@60500: subsection \Basic operations\ wenzelm@60500: wenzelm@60500: subsubsection \Union\ wenzelm@10249: haftmann@34943: lemma count_union [simp]: "count (M + N) a = count M a + count N a" bulwahn@47429: by (simp add: plus_multiset.rep_eq) wenzelm@10249: wenzelm@10249: wenzelm@60500: subsubsection \Difference\ wenzelm@10249: haftmann@49388: instantiation multiset :: (type) comm_monoid_diff haftmann@34943: begin haftmann@34943: haftmann@49388: instance wenzelm@60678: by (standard; transfer; simp add: fun_eq_iff) haftmann@34943: haftmann@34943: end haftmann@34943: haftmann@34943: lemma count_diff [simp]: "count (M - N) a = count M a - count N a" bulwahn@47429: by (simp add: minus_multiset.rep_eq) haftmann@34943: wenzelm@17161: lemma diff_empty [simp]: "M - {#} = M \ {#} - M = {#}" haftmann@52289: by rule (fact Groups.diff_zero, fact Groups.zero_diff) nipkow@36903: nipkow@36903: lemma diff_cancel[simp]: "A - A = {#}" haftmann@52289: by (fact Groups.diff_cancel) wenzelm@10249: nipkow@36903: lemma diff_union_cancelR [simp]: "M + N - N = (M::'a multiset)" haftmann@52289: by (fact add_diff_cancel_right') wenzelm@10249: nipkow@36903: lemma diff_union_cancelL [simp]: "N + M - N = (M::'a multiset)" haftmann@52289: by (fact add_diff_cancel_left') haftmann@34943: haftmann@52289: lemma diff_right_commute: wenzelm@60606: fixes M N Q :: "'a multiset" wenzelm@60606: shows "M - N - Q = M - Q - N" haftmann@52289: by (fact diff_right_commute) haftmann@52289: haftmann@52289: lemma diff_add: wenzelm@60606: fixes M N Q :: "'a multiset" wenzelm@60606: shows "M - (N + Q) = M - N - Q" haftmann@52289: by (rule sym) (fact diff_diff_add) blanchet@58425: wenzelm@60606: lemma insert_DiffM: "x \# M \ {#x#} + (M - {#x#}) = M" nipkow@39302: by (clarsimp simp: multiset_eq_iff) haftmann@34943: wenzelm@60606: lemma insert_DiffM2 [simp]: "x \# M \ M - {#x#} + {#x#} = M" nipkow@39302: by (clarsimp simp: multiset_eq_iff) haftmann@34943: wenzelm@60606: lemma diff_union_swap: "a \ b \ M - {#a#} + {#b#} = M + {#b#} - {#a#}" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: wenzelm@60606: lemma diff_union_single_conv: "a \# J \ I + J - {#a#} = I + (J - {#a#})" nipkow@39302: by (simp add: multiset_eq_iff) bulwahn@26143: wenzelm@10249: wenzelm@60500: subsubsection \Equality of multisets\ haftmann@34943: haftmann@34943: lemma single_not_empty [simp]: "{#a#} \ {#} \ {#} \ {#a#}" nipkow@39302: by (simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma single_eq_single [simp]: "{#a#} = {#b#} \ a = b" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma union_eq_empty [iff]: "M + N = {#} \ M = {#} \ N = {#}" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma empty_eq_union [iff]: "{#} = M + N \ M = {#} \ N = {#}" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma multi_self_add_other_not_self [simp]: "M = M + {#x#} \ False" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: wenzelm@60606: lemma diff_single_trivial: "\ x \# M \ M - {#x#} = M" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: wenzelm@60606: lemma diff_single_eq_union: "x \# M \ M - {#x#} = N \ M = N + {#x#}" haftmann@34943: by auto haftmann@34943: wenzelm@60606: lemma union_single_eq_diff: "M + {#x#} = N \ M = N - {#x#}" haftmann@34943: by (auto dest: sym) haftmann@34943: wenzelm@60606: lemma union_single_eq_member: "M + {#x#} = N \ x \# N" haftmann@34943: by auto haftmann@34943: wenzelm@60606: lemma union_is_single: "M + N = {#a#} \ M = {#a#} \ N={#} \ M = {#} \ N = {#a#}" wenzelm@60606: (is "?lhs = ?rhs") wenzelm@46730: proof wenzelm@60606: show ?lhs if ?rhs using that by auto wenzelm@60606: show ?rhs if ?lhs wenzelm@60606: using that by (simp add: multiset_eq_iff split: if_splits) (metis add_is_1) haftmann@34943: qed haftmann@34943: wenzelm@60606: lemma single_is_union: "{#a#} = M + N \ {#a#} = M \ N = {#} \ M = {#} \ {#a#} = N" haftmann@34943: by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single) haftmann@34943: haftmann@34943: lemma add_eq_conv_diff: wenzelm@60606: "M + {#a#} = N + {#b#} \ M = N \ a = b \ M = N - {#a#} + {#b#} \ N = M - {#b#} + {#a#}" wenzelm@60606: (is "?lhs \ ?rhs") nipkow@44890: (* shorter: by (simp add: multiset_eq_iff) fastforce *) haftmann@34943: proof wenzelm@60606: show ?lhs if ?rhs wenzelm@60606: using that wenzelm@60606: by (auto simp add: add.assoc add.commute [of "{#b#}"]) wenzelm@60606: (drule sym, simp add: add.assoc [symmetric]) wenzelm@60606: show ?rhs if ?lhs haftmann@34943: proof (cases "a = b") wenzelm@60500: case True with \?lhs\ show ?thesis by simp haftmann@34943: next haftmann@34943: case False wenzelm@60500: from \?lhs\ have "a \# N + {#b#}" by (rule union_single_eq_member) haftmann@34943: with False have "a \# N" by auto wenzelm@60500: moreover from \?lhs\ have "M = N + {#b#} - {#a#}" by (rule union_single_eq_diff) haftmann@34943: moreover note False haftmann@34943: ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"] diff_union_swap) haftmann@34943: qed haftmann@34943: qed haftmann@34943: blanchet@58425: lemma insert_noteq_member: haftmann@34943: assumes BC: "B + {#b#} = C + {#c#}" haftmann@34943: and bnotc: "b \ c" haftmann@34943: shows "c \# B" haftmann@34943: proof - haftmann@34943: have "c \# C + {#c#}" by simp haftmann@34943: have nc: "\ c \# {#b#}" using bnotc by simp haftmann@34943: then have "c \# B + {#b#}" using BC by simp haftmann@34943: then show "c \# B" using nc by simp haftmann@34943: qed haftmann@34943: haftmann@34943: lemma add_eq_conv_ex: haftmann@34943: "(M + {#a#} = N + {#b#}) = haftmann@34943: (M = N \ a = b \ (\K. M = K + {#b#} \ N = K + {#a#}))" haftmann@34943: by (auto simp add: add_eq_conv_diff) haftmann@34943: wenzelm@60606: lemma multi_member_split: "x \# M \ \A. M = A + {#x#}" wenzelm@60678: by (rule exI [where x = "M - {#x#}"]) simp haftmann@51600: blanchet@58425: lemma multiset_add_sub_el_shuffle: wenzelm@60606: assumes "c \# B" wenzelm@60606: and "b \ c" haftmann@58098: shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}" haftmann@58098: proof - wenzelm@60500: from \c \# B\ obtain A where B: "B = A + {#c#}" haftmann@58098: by (blast dest: multi_member_split) haftmann@58098: have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp blanchet@58425: then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" haftmann@58098: by (simp add: ac_simps) haftmann@58098: then show ?thesis using B by simp haftmann@58098: qed haftmann@58098: haftmann@34943: wenzelm@60500: subsubsection \Pointwise ordering induced by count\ haftmann@34943: Mathias@60397: definition subseteq_mset :: "'a multiset \ 'a multiset \ bool" (infix "<=#" 50) where Mathias@60400: "subseteq_mset A B = (\a. count A a \ count B a)" Mathias@60397: Mathias@60397: definition subset_mset :: "'a multiset \ 'a multiset \ bool" (infix "<#" 50) where Mathias@60400: "subset_mset A B = (A <=# B \ A \ B)" Mathias@60397: Mathias@60397: notation subseteq_mset (infix "\#" 50) Mathias@60397: notation (xsymbols) subseteq_mset (infix "\#" 50) Mathias@60397: Mathias@60397: notation (xsymbols) subset_mset (infix "\#" 50) Mathias@60397: wenzelm@60606: interpretation subset_mset: ordered_ab_semigroup_add_imp_le "op +" "op -" "op \#" "op \#" wenzelm@60678: by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym) blanchet@59986: wenzelm@60606: lemma mset_less_eqI: "(\x. count A x \ count B x) \ A \# B" Mathias@60397: by (simp add: subseteq_mset_def) haftmann@34943: wenzelm@60606: lemma mset_le_exists_conv: "(A::'a multiset) \# B \ (\C. B = A + C)" wenzelm@60678: unfolding subseteq_mset_def wenzelm@60678: apply (rule iffI) wenzelm@60678: apply (rule exI [where x = "B - A"]) wenzelm@60678: apply (auto intro: multiset_eq_iff [THEN iffD2]) wenzelm@60678: done haftmann@34943: Mathias@60397: interpretation subset_mset: ordered_cancel_comm_monoid_diff "op +" "op -" 0 "op \#" "op <#" wenzelm@60678: by standard (simp, fact mset_le_exists_conv) haftmann@52289: wenzelm@60606: lemma mset_le_mono_add_right_cancel [simp]: "(A::'a multiset) + C \# B + C \ A \# B" Mathias@60397: by (fact subset_mset.add_le_cancel_right) haftmann@34943: wenzelm@60606: lemma mset_le_mono_add_left_cancel [simp]: "C + (A::'a multiset) \# C + B \ A \# B" Mathias@60397: by (fact subset_mset.add_le_cancel_left) haftmann@35268: wenzelm@60606: lemma mset_le_mono_add: "(A::'a multiset) \# B \ C \# D \ A + C \# B + D" Mathias@60397: by (fact subset_mset.add_mono) haftmann@34943: wenzelm@60606: lemma mset_le_add_left [simp]: "(A::'a multiset) \# A + B" Mathias@60397: unfolding subseteq_mset_def by auto haftmann@35268: wenzelm@60606: lemma mset_le_add_right [simp]: "B \# (A::'a multiset) + B" Mathias@60397: unfolding subseteq_mset_def by auto haftmann@34943: wenzelm@60607: lemma mset_le_single: "a \# B \ {#a#} \# B" Mathias@60397: by (simp add: subseteq_mset_def) haftmann@34943: haftmann@35268: lemma multiset_diff_union_assoc: wenzelm@60606: fixes A B C D :: "'a multiset" wenzelm@60606: shows "C \# B \ A + B - C = A + (B - C)" Mathias@60397: by (simp add: subset_mset.diff_add_assoc) haftmann@34943: haftmann@34943: lemma mset_le_multiset_union_diff_commute: wenzelm@60606: fixes A B C D :: "'a multiset" wenzelm@60606: shows "B \# A \ A - B + C = A + C - B" Mathias@60397: by (simp add: subset_mset.add_diff_assoc2) Mathias@60397: Mathias@60397: lemma diff_le_self[simp]: "(M::'a multiset) - N \# M" Mathias@60397: by(simp add: subseteq_mset_def) Mathias@60397: Mathias@60397: lemma mset_lessD: "A <# B \ x \# A \ x \# B" Mathias@60397: apply (clarsimp simp: subset_mset_def subseteq_mset_def) wenzelm@60678: apply (erule allE [where x = x]) haftmann@34943: apply auto haftmann@34943: done haftmann@34943: Mathias@60397: lemma mset_leD: "A \# B \ x \# A \ x \# B" Mathias@60397: apply (clarsimp simp: subset_mset_def subseteq_mset_def) wenzelm@60678: apply (erule allE [where x = x]) haftmann@34943: apply auto haftmann@34943: done blanchet@58425: Mathias@60397: lemma mset_less_insertD: "(A + {#x#} <# B) \ (x \# B \ A <# B)" haftmann@34943: apply (rule conjI) haftmann@34943: apply (simp add: mset_lessD) Mathias@60397: apply (clarsimp simp: subset_mset_def subseteq_mset_def) haftmann@34943: apply safe haftmann@34943: apply (erule_tac x = a in allE) haftmann@34943: apply (auto split: split_if_asm) haftmann@34943: done haftmann@34943: Mathias@60397: lemma mset_le_insertD: "(A + {#x#} \# B) \ (x \# B \ A \# B)" haftmann@34943: apply (rule conjI) haftmann@34943: apply (simp add: mset_leD) Mathias@60397: apply (force simp: subset_mset_def subseteq_mset_def split: split_if_asm) haftmann@34943: done haftmann@34943: Mathias@60397: lemma mset_less_of_empty[simp]: "A <# {#} \ False" Mathias@60397: by (auto simp add: subseteq_mset_def subset_mset_def multiset_eq_iff) Mathias@60397: Mathias@60397: lemma empty_le[simp]: "{#} \# A" nipkow@55808: unfolding mset_le_exists_conv by auto nipkow@55808: Mathias@60397: lemma le_empty[simp]: "(M \# {#}) = (M = {#})" nipkow@55808: unfolding mset_le_exists_conv by auto nipkow@55808: Mathias@60397: lemma multi_psub_of_add_self[simp]: "A <# A + {#x#}" Mathias@60397: by (auto simp: subset_mset_def subseteq_mset_def) Mathias@60397: Mathias@60397: lemma multi_psub_self[simp]: "(A::'a multiset) <# A = False" haftmann@35268: by simp haftmann@34943: Mathias@60397: lemma mset_less_add_bothsides: "N + {#x#} <# M + {#x#} \ N <# M" Mathias@60397: by (fact subset_mset.add_less_imp_less_right) haftmann@35268: wenzelm@60606: lemma mset_less_empty_nonempty: "{#} <# S \ S \ {#}" Mathias@60397: by (auto simp: subset_mset_def subseteq_mset_def) haftmann@35268: wenzelm@60606: lemma mset_less_diff_self: "c \# B \ B - {#c#} <# B" Mathias@60397: by (auto simp: subset_mset_def subseteq_mset_def multiset_eq_iff) haftmann@35268: haftmann@35268: wenzelm@60500: subsubsection \Intersection\ haftmann@35268: Mathias@60397: definition inf_subset_mset :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "#\" 70) where Mathias@60397: multiset_inter_def: "inf_subset_mset A B = A - (A - B)" Mathias@60397: Mathias@60397: interpretation subset_mset: semilattice_inf inf_subset_mset "op \#" "op <#" wenzelm@46921: proof - wenzelm@60678: have [simp]: "m \ n \ m \ q \ m \ n - (n - q)" for m n q :: nat wenzelm@60678: by arith wenzelm@60678: show "class.semilattice_inf op #\ op \# op <#" wenzelm@60678: by standard (auto simp add: multiset_inter_def subseteq_mset_def) haftmann@35268: qed haftmann@35268: haftmann@34943: haftmann@41069: lemma multiset_inter_count [simp]: wenzelm@60606: fixes A B :: "'a multiset" wenzelm@60606: shows "count (A #\ B) x = min (count A x) (count B x)" bulwahn@47429: by (simp add: multiset_inter_def) haftmann@35268: haftmann@35268: lemma multiset_inter_single: "a \ b \ {#a#} #\ {#b#} = {#}" wenzelm@46730: by (rule multiset_eqI) auto haftmann@34943: haftmann@35268: lemma multiset_union_diff_commute: haftmann@35268: assumes "B #\ C = {#}" haftmann@35268: shows "A + B - C = A - C + B" nipkow@39302: proof (rule multiset_eqI) haftmann@35268: fix x haftmann@35268: from assms have "min (count B x) (count C x) = 0" wenzelm@46730: by (auto simp add: multiset_eq_iff) haftmann@35268: then have "count B x = 0 \ count C x = 0" haftmann@35268: by auto haftmann@35268: then show "count (A + B - C) x = count (A - C + B) x" haftmann@35268: by auto haftmann@35268: qed haftmann@35268: wenzelm@60606: lemma empty_inter [simp]: "{#} #\ M = {#}" haftmann@51600: by (simp add: multiset_eq_iff) haftmann@51600: wenzelm@60606: lemma inter_empty [simp]: "M #\ {#} = {#}" haftmann@51600: by (simp add: multiset_eq_iff) haftmann@51600: wenzelm@60606: lemma inter_add_left1: "\ x \# N \ (M + {#x#}) #\ N = M #\ N" haftmann@51600: by (simp add: multiset_eq_iff) haftmann@51600: wenzelm@60606: lemma inter_add_left2: "x \# N \ (M + {#x#}) #\ N = (M #\ (N - {#x#})) + {#x#}" haftmann@51600: by (simp add: multiset_eq_iff) haftmann@51600: wenzelm@60606: lemma inter_add_right1: "\ x \# N \ N #\ (M + {#x#}) = N #\ M" haftmann@51600: by (simp add: multiset_eq_iff) haftmann@51600: wenzelm@60606: lemma inter_add_right2: "x \# N \ N #\ (M + {#x#}) = ((N - {#x#}) #\ M) + {#x#}" haftmann@51600: by (simp add: multiset_eq_iff) haftmann@51600: haftmann@35268: wenzelm@60500: subsubsection \Bounded union\ wenzelm@60678: wenzelm@60678: definition sup_subset_mset :: "'a multiset \ 'a multiset \ 'a multiset"(infixl "#\" 70) wenzelm@60678: where "sup_subset_mset A B = A + (B - A)" Mathias@60397: Mathias@60397: interpretation subset_mset: semilattice_sup sup_subset_mset "op \#" "op <#" haftmann@51623: proof - wenzelm@60678: have [simp]: "m \ n \ q \ n \ m + (q - m) \ n" for m n q :: nat wenzelm@60678: by arith Mathias@60397: show "class.semilattice_sup op #\ op \# op <#" wenzelm@60678: by standard (auto simp add: sup_subset_mset_def subseteq_mset_def) haftmann@51623: qed haftmann@51623: wenzelm@60606: lemma sup_subset_mset_count [simp]: "count (A #\ B) x = max (count A x) (count B x)" Mathias@60397: by (simp add: sup_subset_mset_def) haftmann@51623: wenzelm@60606: lemma empty_sup [simp]: "{#} #\ M = M" haftmann@51623: by (simp add: multiset_eq_iff) haftmann@51623: wenzelm@60606: lemma sup_empty [simp]: "M #\ {#} = M" haftmann@51623: by (simp add: multiset_eq_iff) haftmann@51623: wenzelm@60606: lemma sup_add_left1: "\ x \# N \ (M + {#x#}) #\ N = (M #\ N) + {#x#}" haftmann@51623: by (simp add: multiset_eq_iff) haftmann@51623: wenzelm@60606: lemma sup_add_left2: "x \# N \ (M + {#x#}) #\ N = (M #\ (N - {#x#})) + {#x#}" haftmann@51623: by (simp add: multiset_eq_iff) haftmann@51623: wenzelm@60606: lemma sup_add_right1: "\ x \# N \ N #\ (M + {#x#}) = (N #\ M) + {#x#}" haftmann@51623: by (simp add: multiset_eq_iff) haftmann@51623: wenzelm@60606: lemma sup_add_right2: "x \# N \ N #\ (M + {#x#}) = ((N - {#x#}) #\ M) + {#x#}" haftmann@51623: by (simp add: multiset_eq_iff) haftmann@51623: wenzelm@60500: subsubsection \Subset is an order\ Mathias@60397: interpretation subset_mset: order "op \#" "op <#" by unfold_locales auto haftmann@51623: wenzelm@60500: subsubsection \Filter (with comprehension syntax)\ wenzelm@60500: wenzelm@60500: text \Multiset comprehension\ haftmann@41069: nipkow@59998: lift_definition filter_mset :: "('a \ bool) \ 'a multiset \ 'a multiset" nipkow@59998: is "\P M. \x. if P x then M x else 0" bulwahn@47429: by (rule filter_preserves_multiset) haftmann@35268: wenzelm@60606: lemma count_filter_mset [simp]: "count (filter_mset P M) a = (if P a then count M a else 0)" nipkow@59998: by (simp add: filter_mset.rep_eq) nipkow@59998: wenzelm@60606: lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}" nipkow@59998: by (rule multiset_eqI) simp nipkow@59998: wenzelm@60606: lemma filter_single_mset [simp]: "filter_mset P {#x#} = (if P x then {#x#} else {#})" nipkow@39302: by (rule multiset_eqI) simp haftmann@35268: wenzelm@60606: lemma filter_union_mset [simp]: "filter_mset P (M + N) = filter_mset P M + filter_mset P N" haftmann@41069: by (rule multiset_eqI) simp haftmann@41069: wenzelm@60606: lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N" nipkow@39302: by (rule multiset_eqI) simp haftmann@35268: wenzelm@60606: lemma filter_inter_mset [simp]: "filter_mset P (M #\ N) = filter_mset P M #\ filter_mset P N" haftmann@41069: by (rule multiset_eqI) simp haftmann@41069: Mathias@60397: lemma multiset_filter_subset[simp]: "filter_mset f M \# M" Mathias@60397: by (simp add: mset_less_eqI) Mathias@60397: wenzelm@60606: lemma multiset_filter_mono: wenzelm@60606: assumes "A \# B" Mathias@60397: shows "filter_mset f A \# filter_mset f B" blanchet@58035: proof - blanchet@58035: from assms[unfolded mset_le_exists_conv] blanchet@58035: obtain C where B: "B = A + C" by auto blanchet@58035: show ?thesis unfolding B by auto blanchet@58035: qed blanchet@58035: haftmann@41069: syntax haftmann@41069: "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{# _ :# _./ _#})") haftmann@41069: syntax (xsymbol) haftmann@41069: "_MCollect" :: "pttrn \ 'a multiset \ bool \ 'a multiset" ("(1{# _ \# _./ _#})") haftmann@41069: translations nipkow@59998: "{#x \# M. P#}" == "CONST filter_mset (\x. P) M" haftmann@41069: wenzelm@10249: wenzelm@60500: subsubsection \Set of elements\ wenzelm@10249: wenzelm@60606: definition set_mset :: "'a multiset \ 'a set" wenzelm@60607: where "set_mset M = {x. x \# M}" nipkow@60495: nipkow@60495: lemma set_mset_empty [simp]: "set_mset {#} = {}" nipkow@60495: by (simp add: set_mset_def) nipkow@60495: nipkow@60495: lemma set_mset_single [simp]: "set_mset {#b#} = {b}" nipkow@60495: by (simp add: set_mset_def) nipkow@60495: nipkow@60495: lemma set_mset_union [simp]: "set_mset (M + N) = set_mset M \ set_mset N" nipkow@60495: by (auto simp add: set_mset_def) nipkow@60495: nipkow@60495: lemma set_mset_eq_empty_iff [simp]: "(set_mset M = {}) = (M = {#})" nipkow@60495: by (auto simp add: set_mset_def multiset_eq_iff) nipkow@60495: wenzelm@60607: lemma mem_set_mset_iff [simp]: "(x \ set_mset M) = (x \# M)" nipkow@60495: by (auto simp add: set_mset_def) nipkow@60495: wenzelm@60607: lemma set_mset_filter [simp]: "set_mset {# x\#M. P x #} = set_mset M \ {x. P x}" nipkow@60495: by (auto simp add: set_mset_def) nipkow@60495: nipkow@60495: lemma finite_set_mset [iff]: "finite (set_mset M)" nipkow@60495: using count [of M] by (simp add: multiset_def set_mset_def) haftmann@34943: wenzelm@60607: lemma finite_Collect_mem [iff]: "finite {x. x \# M}" nipkow@60495: unfolding set_mset_def[symmetric] by simp nipkow@60495: nipkow@60495: lemma set_mset_mono: "A \# B \ set_mset A \ set_mset B" nipkow@60495: by (metis mset_leD subsetI mem_set_mset_iff) nipkow@60495: nipkow@60495: lemma ball_set_mset_iff: "(\x \ set_mset M. P x) \ (\x. x \# M \ P x)" blanchet@59813: by auto blanchet@59813: blanchet@59813: wenzelm@60500: subsubsection \Size\ wenzelm@10249: blanchet@56656: definition wcount where "wcount f M = (\x. count M x * Suc (f x))" blanchet@56656: blanchet@56656: lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" blanchet@56656: by (auto simp: wcount_def add_mult_distrib) blanchet@56656: blanchet@56656: definition size_multiset :: "('a \ nat) \ 'a multiset \ nat" where nipkow@60495: "size_multiset f M = setsum (wcount f M) (set_mset M)" blanchet@56656: blanchet@56656: lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] blanchet@56656: wenzelm@60606: instantiation multiset :: (type) size wenzelm@60606: begin wenzelm@60606: blanchet@56656: definition size_multiset where blanchet@56656: size_multiset_overloaded_def: "size_multiset = Multiset.size_multiset (\_. 0)" haftmann@34943: instance .. wenzelm@60606: haftmann@34943: end haftmann@34943: blanchet@56656: lemmas size_multiset_overloaded_eq = blanchet@56656: size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] blanchet@56656: blanchet@56656: lemma size_multiset_empty [simp]: "size_multiset f {#} = 0" blanchet@56656: by (simp add: size_multiset_def) blanchet@56656: haftmann@28708: lemma size_empty [simp]: "size {#} = 0" blanchet@56656: by (simp add: size_multiset_overloaded_def) blanchet@56656: blanchet@56656: lemma size_multiset_single [simp]: "size_multiset f {#b#} = Suc (f b)" blanchet@56656: by (simp add: size_multiset_eq) wenzelm@10249: haftmann@28708: lemma size_single [simp]: "size {#b#} = 1" blanchet@56656: by (simp add: size_multiset_overloaded_def) blanchet@56656: blanchet@56656: lemma setsum_wcount_Int: nipkow@60495: "finite A \ setsum (wcount f N) (A \ set_mset N) = setsum (wcount f N) A" nipkow@26178: apply (induct rule: finite_induct) nipkow@26178: apply simp nipkow@60495: apply (simp add: Int_insert_left set_mset_def wcount_def) blanchet@56656: done blanchet@56656: blanchet@56656: lemma size_multiset_union [simp]: blanchet@56656: "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" haftmann@57418: apply (simp add: size_multiset_def setsum_Un_nat setsum.distrib setsum_wcount_Int wcount_union) blanchet@56656: apply (subst Int_commute) blanchet@56656: apply (simp add: setsum_wcount_Int) nipkow@26178: done wenzelm@10249: haftmann@28708: lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" blanchet@56656: by (auto simp add: size_multiset_overloaded_def) blanchet@56656: blanchet@56656: lemma size_multiset_eq_0_iff_empty [iff]: "(size_multiset f M = 0) = (M = {#})" blanchet@56656: by (auto simp add: size_multiset_eq multiset_eq_iff) wenzelm@10249: wenzelm@17161: lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" blanchet@56656: by (auto simp add: size_multiset_overloaded_def) 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@60607: lemma size_eq_Suc_imp_elem: "size M = Suc n \ \a. a \# M" blanchet@56656: apply (unfold size_multiset_overloaded_eq) nipkow@26178: apply (drule setsum_SucD) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: haftmann@34943: lemma size_eq_Suc_imp_eq_union: haftmann@34943: assumes "size M = Suc n" haftmann@34943: shows "\a N. M = N + {#a#}" haftmann@34943: proof - haftmann@34943: from assms obtain a where "a \# M" haftmann@34943: by (erule size_eq_Suc_imp_elem [THEN exE]) haftmann@34943: then have "M = M - {#a#} + {#a#}" by simp haftmann@34943: then show ?thesis by blast nipkow@23611: qed kleing@15869: wenzelm@60606: lemma size_mset_mono: wenzelm@60606: fixes A B :: "'a multiset" wenzelm@60606: assumes "A \# B" wenzelm@60606: shows "size A \ size B" nipkow@59949: proof - nipkow@59949: from assms[unfolded mset_le_exists_conv] nipkow@59949: obtain C where B: "B = A + C" by auto wenzelm@60606: show ?thesis unfolding B by (induct C) auto nipkow@59949: qed nipkow@59949: nipkow@59998: lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \ size M" nipkow@59949: by (rule size_mset_mono[OF multiset_filter_subset]) nipkow@59949: nipkow@59949: lemma size_Diff_submset: Mathias@60397: "M \# M' \ size (M' - M) = size M' - size(M::'a multiset)" nipkow@59949: by (metis add_diff_cancel_left' size_union mset_le_exists_conv) nipkow@26016: wenzelm@60500: subsection \Induction and case splits\ wenzelm@10249: wenzelm@18258: theorem multiset_induct [case_names empty add, induct type: multiset]: huffman@48009: assumes empty: "P {#}" huffman@48009: assumes add: "\M x. P M \ P (M + {#x#})" huffman@48009: shows "P M" huffman@48009: proof (induct n \ "size M" arbitrary: M) huffman@48009: case 0 thus "P M" by (simp add: empty) huffman@48009: next huffman@48009: case (Suc k) huffman@48009: obtain N x where "M = N + {#x#}" wenzelm@60500: using \Suc k = size M\ [symmetric] huffman@48009: using size_eq_Suc_imp_eq_union by fast huffman@48009: with Suc add show "P M" by simp 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: wenzelm@55913: lemma multiset_cases [cases type]: wenzelm@55913: obtains (empty) "M = {#}" wenzelm@55913: | (add) N x where "M = N + {#x#}" wenzelm@55913: using assms by (induct M) simp_all kleing@25610: haftmann@34943: lemma multi_drop_mem_not_eq: "c \# B \ B - {#c#} \ B" haftmann@34943: by (cases "B = {#}") (auto dest: multi_member_split) haftmann@34943: wenzelm@60607: lemma multiset_partition: "M = {# x\#M. P x #} + {# x\#M. \ P x #}" nipkow@39302: apply (subst multiset_eq_iff) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: Mathias@60397: lemma mset_less_size: "(A::'a multiset) <# B \ size A < size B" haftmann@34943: proof (induct A arbitrary: B) haftmann@34943: case (empty M) haftmann@34943: then have "M \ {#}" by (simp add: mset_less_empty_nonempty) blanchet@58425: then obtain M' x where "M = M' + {#x#}" haftmann@34943: by (blast dest: multi_nonempty_split) haftmann@34943: then show ?case by simp haftmann@34943: next haftmann@34943: case (add S x T) Mathias@60397: have IH: "\B. S <# B \ size S < size B" by fact Mathias@60397: have SxsubT: "S + {#x#} <# T" by fact Mathias@60397: then have "x \# T" and "S <# T" by (auto dest: mset_less_insertD) blanchet@58425: then obtain T' where T: "T = T' + {#x#}" haftmann@34943: by (blast dest: multi_member_split) Mathias@60397: then have "S <# T'" using SxsubT haftmann@34943: by (blast intro: mset_less_add_bothsides) haftmann@34943: then have "size S < size T'" using IH by simp haftmann@34943: then show ?case using T by simp haftmann@34943: qed haftmann@34943: haftmann@34943: nipkow@59949: lemma size_1_singleton_mset: "size M = 1 \ \a. M = {#a#}" nipkow@59949: by (cases M) auto nipkow@59949: wenzelm@60500: subsubsection \Strong induction and subset induction for multisets\ wenzelm@60500: wenzelm@60500: text \Well-foundedness of strict subset relation\ haftmann@58098: Mathias@60397: lemma wf_less_mset_rel: "wf {(M, N :: 'a multiset). M <# N}" haftmann@34943: apply (rule wf_measure [THEN wf_subset, where f1=size]) haftmann@34943: apply (clarsimp simp: measure_def inv_image_def mset_less_size) haftmann@34943: done haftmann@34943: haftmann@34943: lemma full_multiset_induct [case_names less]: Mathias@60397: assumes ih: "\B. \(A::'a multiset). A <# B \ P A \ P B" haftmann@34943: shows "P B" haftmann@58098: apply (rule wf_less_mset_rel [THEN wf_induct]) haftmann@58098: apply (rule ih, auto) haftmann@34943: done haftmann@34943: haftmann@34943: lemma multi_subset_induct [consumes 2, case_names empty add]: wenzelm@60606: assumes "F \# A" wenzelm@60606: and empty: "P {#}" wenzelm@60606: and insert: "\a F. a \# A \ P F \ P (F + {#a#})" wenzelm@60606: shows "P F" haftmann@34943: proof - wenzelm@60500: from \F \# A\ haftmann@34943: show ?thesis haftmann@34943: proof (induct F) haftmann@34943: show "P {#}" by fact haftmann@34943: next haftmann@34943: fix x F Mathias@60397: assume P: "F \# A \ P F" and i: "F + {#x#} \# A" haftmann@34943: show "P (F + {#x#})" haftmann@34943: proof (rule insert) haftmann@34943: from i show "x \# A" by (auto dest: mset_le_insertD) Mathias@60397: from i have "F \# A" by (auto dest: mset_le_insertD) haftmann@34943: with P show "P F" . haftmann@34943: qed haftmann@34943: qed haftmann@34943: qed wenzelm@26145: wenzelm@17161: wenzelm@60500: subsection \The fold combinator\ huffman@48023: nipkow@59998: definition fold_mset :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b" huffman@48023: where nipkow@60495: "fold_mset f s M = Finite_Set.fold (\x. f x ^^ count M x) s (set_mset M)" huffman@48023: wenzelm@60606: lemma fold_mset_empty [simp]: "fold_mset f s {#} = s" nipkow@59998: by (simp add: fold_mset_def) huffman@48023: huffman@48023: context comp_fun_commute huffman@48023: begin huffman@48023: wenzelm@60606: lemma fold_mset_insert: "fold_mset f s (M + {#x#}) = f x (fold_mset f s M)" haftmann@49822: proof - haftmann@49822: interpret mset: comp_fun_commute "\y. f y ^^ count M y" haftmann@49822: by (fact comp_fun_commute_funpow) haftmann@49822: interpret mset_union: comp_fun_commute "\y. f y ^^ count (M + {#x#}) y" haftmann@49822: by (fact comp_fun_commute_funpow) haftmann@49822: show ?thesis nipkow@60495: proof (cases "x \ set_mset M") haftmann@49822: case False haftmann@49822: then have *: "count (M + {#x#}) x = 1" by simp nipkow@60495: from False have "Finite_Set.fold (\y. f y ^^ count (M + {#x#}) y) s (set_mset M) = nipkow@60495: Finite_Set.fold (\y. f y ^^ count M y) s (set_mset M)" haftmann@49822: by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) haftmann@49822: with False * show ?thesis nipkow@59998: by (simp add: fold_mset_def del: count_union) huffman@48023: next haftmann@49822: case True nipkow@60495: def N \ "set_mset M - {x}" nipkow@60495: from N_def True have *: "set_mset M = insert x N" "x \ N" "finite N" by auto haftmann@49822: then have "Finite_Set.fold (\y. f y ^^ count (M + {#x#}) y) s N = haftmann@49822: Finite_Set.fold (\y. f y ^^ count M y) s N" haftmann@49822: by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) nipkow@59998: with * show ?thesis by (simp add: fold_mset_def del: count_union) simp huffman@48023: qed huffman@48023: qed huffman@48023: wenzelm@60606: corollary fold_mset_single [simp]: "fold_mset f s {#x#} = f x s" haftmann@49822: proof - nipkow@59998: have "fold_mset f s ({#} + {#x#}) = f x s" by (simp only: fold_mset_insert) simp haftmann@49822: then show ?thesis by simp haftmann@49822: qed huffman@48023: wenzelm@60606: lemma fold_mset_fun_left_comm: "f x (fold_mset f s M) = fold_mset f (f x s) M" haftmann@49822: by (induct M) (simp_all add: fold_mset_insert fun_left_comm) huffman@48023: wenzelm@60606: lemma fold_mset_union [simp]: "fold_mset f s (M + N) = fold_mset f (fold_mset f s M) N" haftmann@49822: proof (induct M) huffman@48023: case empty then show ?case by simp huffman@48023: next haftmann@49822: case (add M x) haftmann@49822: have "M + {#x#} + N = (M + N) + {#x#}" haftmann@57514: by (simp add: ac_simps) haftmann@51548: with add show ?case by (simp add: fold_mset_insert fold_mset_fun_left_comm) huffman@48023: qed huffman@48023: huffman@48023: lemma fold_mset_fusion: huffman@48023: assumes "comp_fun_commute g" wenzelm@60606: and *: "\x y. h (g x y) = f x (h y)" wenzelm@60606: shows "h (fold_mset g w A) = fold_mset f (h w) A" huffman@48023: proof - huffman@48023: interpret comp_fun_commute g by (fact assms) wenzelm@60606: from * show ?thesis by (induct A) auto huffman@48023: qed huffman@48023: huffman@48023: end huffman@48023: wenzelm@60500: text \ huffman@48023: A note on code generation: When defining some function containing a nipkow@59998: subterm @{term "fold_mset F"}, code generation is not automatic. When wenzelm@61585: interpreting locale \left_commutative\ with \F\, the nipkow@59998: would be code thms for @{const fold_mset} become thms like wenzelm@61585: @{term "fold_mset F z {#} = z"} where \F\ is not a pattern but huffman@48023: contains defined symbols, i.e.\ is not a code thm. Hence a separate wenzelm@61585: constant with its own code thms needs to be introduced for \F\. See the image operator below. wenzelm@60500: \ wenzelm@60500: wenzelm@60500: wenzelm@60500: subsection \Image\ huffman@48023: huffman@48023: definition image_mset :: "('a \ 'b) \ 'a multiset \ 'b multiset" where wenzelm@60607: "image_mset f = fold_mset (plus \ single \ f) {#}" wenzelm@60607: wenzelm@60607: lemma comp_fun_commute_mset_image: "comp_fun_commute (plus \ single \ f)" haftmann@49823: proof haftmann@57514: qed (simp add: ac_simps fun_eq_iff) huffman@48023: huffman@48023: lemma image_mset_empty [simp]: "image_mset f {#} = {#}" haftmann@49823: by (simp add: image_mset_def) huffman@48023: huffman@48023: lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}" haftmann@49823: proof - wenzelm@60607: interpret comp_fun_commute "plus \ single \ f" haftmann@49823: by (fact comp_fun_commute_mset_image) haftmann@49823: show ?thesis by (simp add: image_mset_def) haftmann@49823: qed huffman@48023: wenzelm@60606: lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" haftmann@49823: proof - wenzelm@60607: interpret comp_fun_commute "plus \ single \ f" haftmann@49823: by (fact comp_fun_commute_mset_image) haftmann@57514: show ?thesis by (induct N) (simp_all add: image_mset_def ac_simps) haftmann@49823: qed haftmann@49823: wenzelm@60606: corollary image_mset_insert: "image_mset f (M + {#a#}) = image_mset f M + {#f a#}" haftmann@49823: by simp huffman@48023: wenzelm@60606: lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)" haftmann@49823: by (induct M) simp_all huffman@48040: wenzelm@60606: lemma size_image_mset [simp]: "size (image_mset f M) = size M" haftmann@49823: by (induct M) simp_all huffman@48023: wenzelm@60606: lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" haftmann@49823: by (cases M) auto huffman@48023: huffman@48023: syntax huffman@48023: "_comprehension1_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" huffman@48023: ("({#_/. _ :# _#})") huffman@48023: translations wenzelm@60606: "{#e. x:#M#}" == "CONST image_mset (\x. e) M" huffman@48023: blanchet@59813: syntax (xsymbols) blanchet@59813: "_comprehension2_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" blanchet@59813: ("({#_/. _ \# _#})") blanchet@59813: translations blanchet@59813: "{#e. x \# M#}" == "CONST image_mset (\x. e) M" blanchet@59813: huffman@48023: syntax blanchet@59813: "_comprehension3_mset" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" huffman@48023: ("({#_/ | _ :# _./ _#})") huffman@48023: translations wenzelm@60606: "{#e | x:#M. P#}" \ "{#e. x :# {# x:#M. P#}#}" huffman@48023: blanchet@59813: syntax blanchet@59813: "_comprehension4_mset" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" blanchet@59813: ("({#_/ | _ \# _./ _#})") blanchet@59813: translations wenzelm@60606: "{#e | x\#M. P#}" \ "{#e. x \# {# x\#M. P#}#}" blanchet@59813: wenzelm@60500: text \ wenzelm@60607: This allows to write not just filters like @{term "{#x\#M. x#M #}"} and @{term [source] wenzelm@60607: "{#x+x|x\#M. x#M. x huffman@48023: nipkow@60495: lemma in_image_mset: "y \# {#f x. x \# M#} \ y \ f ` set_mset M" nipkow@60498: by (metis mem_set_mset_iff set_image_mset) blanchet@59813: blanchet@55467: functor image_mset: image_mset huffman@48023: proof - huffman@48023: fix f g show "image_mset f \ image_mset g = image_mset (f \ g)" huffman@48023: proof huffman@48023: fix A huffman@48023: show "(image_mset f \ image_mset g) A = image_mset (f \ g) A" huffman@48023: by (induct A) simp_all huffman@48023: qed huffman@48023: show "image_mset id = id" huffman@48023: proof huffman@48023: fix A huffman@48023: show "image_mset id A = id A" huffman@48023: by (induct A) simp_all huffman@48023: qed huffman@48023: qed huffman@48023: blanchet@59813: declare blanchet@59813: image_mset.id [simp] blanchet@59813: image_mset.identity [simp] blanchet@59813: blanchet@59813: lemma image_mset_id[simp]: "image_mset id x = x" blanchet@59813: unfolding id_def by auto blanchet@59813: blanchet@59813: lemma image_mset_cong: "(\x. x \# M \ f x = g x) \ {#f x. x \# M#} = {#g x. x \# M#}" blanchet@59813: by (induct M) auto blanchet@59813: blanchet@59813: lemma image_mset_cong_pair: blanchet@59813: "(\x y. (x, y) \# M \ f x y = g x y) \ {#f x y. (x, y) \# M#} = {#g x y. (x, y) \# M#}" blanchet@59813: by (metis image_mset_cong split_cong) haftmann@49717: huffman@48023: wenzelm@60500: subsection \Further conversions\ haftmann@34943: nipkow@60515: primrec mset :: "'a list \ 'a multiset" where nipkow@60515: "mset [] = {#}" | nipkow@60515: "mset (a # x) = mset x + {# a #}" haftmann@34943: haftmann@37107: lemma in_multiset_in_set: nipkow@60515: "x \# mset xs \ x \ set xs" haftmann@37107: by (induct xs) simp_all haftmann@37107: nipkow@60515: lemma count_mset: nipkow@60515: "count (mset xs) x = length (filter (\y. x = y) xs)" haftmann@37107: by (induct xs) simp_all haftmann@37107: nipkow@60515: lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])" blanchet@59813: by (induct x) auto haftmann@34943: nipkow@60515: lemma mset_zero_iff_right[simp]: "({#} = mset x) = (x = [])" haftmann@34943: by (induct x) auto haftmann@34943: nipkow@60515: lemma set_mset_mset[simp]: "set_mset (mset x) = set x" haftmann@34943: by (induct x) auto haftmann@34943: wenzelm@60607: lemma mem_set_multiset_eq: "x \ set xs = (x \# mset xs)" haftmann@34943: by (induct xs) auto haftmann@34943: nipkow@60515: lemma size_mset [simp]: "size (mset xs) = length xs" huffman@48012: by (induct xs) simp_all huffman@48012: wenzelm@60606: lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys" haftmann@57514: by (induct xs arbitrary: ys) (auto simp: ac_simps) haftmann@34943: wenzelm@60607: lemma mset_filter: "mset (filter P xs) = {#x \# mset xs. P x #}" haftmann@40303: by (induct xs) simp_all haftmann@40303: nipkow@60515: lemma mset_rev [simp]: nipkow@60515: "mset (rev xs) = mset xs" haftmann@40950: by (induct xs) simp_all haftmann@40950: nipkow@60515: lemma surj_mset: "surj mset" haftmann@34943: apply (unfold surj_def) haftmann@34943: apply (rule allI) haftmann@34943: apply (rule_tac M = y in multiset_induct) haftmann@34943: apply auto haftmann@34943: apply (rule_tac x = "x # xa" in exI) haftmann@34943: apply auto haftmann@34943: done haftmann@34943: nipkow@60515: lemma set_count_greater_0: "set x = {a. count (mset x) a > 0}" haftmann@34943: by (induct x) auto haftmann@34943: haftmann@34943: lemma distinct_count_atmost_1: wenzelm@60606: "distinct x = (\a. count (mset x) a = (if a \ set x then 1 else 0))" wenzelm@60678: apply (induct x, simp, rule iffI, simp_all) wenzelm@60678: subgoal for a b wenzelm@60678: apply (rule conjI) wenzelm@60678: apply (simp_all add: set_mset_mset [symmetric] del: set_mset_mset) wenzelm@60678: apply (erule_tac x = a in allE, simp) wenzelm@60678: apply clarify wenzelm@60678: apply (erule_tac x = aa in allE, simp) wenzelm@60678: done wenzelm@60678: done haftmann@34943: wenzelm@60606: lemma mset_eq_setD: "mset xs = mset ys \ set xs = set ys" nipkow@39302: by (rule) (auto simp add:multiset_eq_iff set_count_greater_0) haftmann@34943: nipkow@60515: lemma set_eq_iff_mset_eq_distinct: haftmann@34943: "distinct x \ distinct y \ nipkow@60515: (set x = set y) = (mset x = mset y)" nipkow@39302: by (auto simp: multiset_eq_iff distinct_count_atmost_1) haftmann@34943: nipkow@60515: lemma set_eq_iff_mset_remdups_eq: nipkow@60515: "(set x = set y) = (mset (remdups x) = mset (remdups y))" haftmann@34943: apply (rule iffI) nipkow@60515: apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1]) haftmann@34943: apply (drule distinct_remdups [THEN distinct_remdups nipkow@60515: [THEN set_eq_iff_mset_eq_distinct [THEN iffD2]]]) haftmann@34943: apply simp haftmann@34943: done haftmann@34943: wenzelm@60606: lemma mset_compl_union [simp]: "mset [x\xs. P x] + mset [x\xs. \P x] = mset xs" haftmann@57514: by (induct xs) (auto simp: ac_simps) haftmann@34943: wenzelm@60607: lemma nth_mem_mset: "i < length ls \ (ls ! i) \# mset ls" wenzelm@60678: proof (induct ls arbitrary: i) wenzelm@60678: case Nil wenzelm@60678: then show ?case by simp wenzelm@60678: next wenzelm@60678: case Cons wenzelm@60678: then show ?case by (cases i) auto wenzelm@60678: qed haftmann@34943: wenzelm@60606: lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}" wenzelm@60678: by (induct xs) (auto simp add: multiset_eq_iff) haftmann@34943: nipkow@60515: lemma mset_eq_length: nipkow@60515: assumes "mset xs = mset ys" haftmann@37107: shows "length xs = length ys" nipkow@60515: using assms by (metis size_mset) nipkow@60515: nipkow@60515: lemma mset_eq_length_filter: nipkow@60515: assumes "mset xs = mset ys" haftmann@39533: shows "length (filter (\x. z = x) xs) = length (filter (\y. z = y) ys)" nipkow@60515: using assms by (metis count_mset) haftmann@39533: haftmann@45989: lemma fold_multiset_equiv: haftmann@45989: assumes f: "\x y. x \ set xs \ y \ set xs \ f x \ f y = f y \ f x" nipkow@60515: and equiv: "mset xs = mset ys" haftmann@49822: shows "List.fold f xs = List.fold f ys" wenzelm@60606: using f equiv [symmetric] wenzelm@46921: proof (induct xs arbitrary: ys) wenzelm@60678: case Nil wenzelm@60678: then show ?case by simp haftmann@45989: next haftmann@45989: case (Cons x xs) wenzelm@60678: then have *: "set ys = set (x # xs)" wenzelm@60678: by (blast dest: mset_eq_setD) blanchet@58425: have "\x y. x \ set ys \ y \ set ys \ f x \ f y = f y \ f x" haftmann@45989: by (rule Cons.prems(1)) (simp_all add: *) wenzelm@60678: moreover from * have "x \ set ys" wenzelm@60678: by simp wenzelm@60678: ultimately have "List.fold f ys = List.fold f (remove1 x ys) \ f x" wenzelm@60678: by (fact fold_remove1_split) wenzelm@60678: moreover from Cons.prems have "List.fold f xs = List.fold f (remove1 x ys)" wenzelm@60678: by (auto intro: Cons.hyps) haftmann@45989: ultimately show ?case by simp haftmann@45989: qed haftmann@45989: wenzelm@60606: lemma mset_insort [simp]: "mset (insort x xs) = mset xs + {#x#}" haftmann@51548: by (induct xs) (simp_all add: ac_simps) haftmann@51548: wenzelm@60606: lemma mset_map: "mset (map f xs) = image_mset f (mset xs)" haftmann@51600: by (induct xs) simp_all haftmann@51600: nipkow@60513: definition mset_set :: "'a set \ 'a multiset" haftmann@51548: where nipkow@60513: "mset_set = folding.F (\x M. {#x#} + M) {#}" nipkow@60513: nipkow@60513: interpretation mset_set!: folding "\x M. {#x#} + M" "{#}" ballarin@61566: rewrites nipkow@60513: "folding.F (\x M. {#x#} + M) {#} = mset_set" haftmann@51548: proof - wenzelm@60678: interpret comp_fun_commute "\x M. {#x#} + M" wenzelm@60678: by standard (simp add: fun_eq_iff ac_simps) wenzelm@60678: show "folding (\x M. {#x#} + M)" wenzelm@60678: by standard (fact comp_fun_commute) nipkow@60513: from mset_set_def show "folding.F (\x M. {#x#} + M) {#} = mset_set" .. haftmann@51548: qed haftmann@51548: nipkow@60513: lemma count_mset_set [simp]: nipkow@60513: "finite A \ x \ A \ count (mset_set A) x = 1" (is "PROP ?P") nipkow@60513: "\ finite A \ count (mset_set A) x = 0" (is "PROP ?Q") nipkow@60513: "x \ A \ count (mset_set A) x = 0" (is "PROP ?R") haftmann@51600: proof - wenzelm@60606: have *: "count (mset_set A) x = 0" if "x \ A" for A wenzelm@60606: proof (cases "finite A") wenzelm@60606: case False then show ?thesis by simp wenzelm@60606: next wenzelm@60606: case True from True \x \ A\ show ?thesis by (induct A) auto wenzelm@60606: qed haftmann@51600: then show "PROP ?P" "PROP ?Q" "PROP ?R" haftmann@51600: by (auto elim!: Set.set_insert) wenzelm@61585: qed \ \TODO: maybe define @{const mset_set} also in terms of @{const Abs_multiset}\ nipkow@60513: nipkow@60513: lemma elem_mset_set[simp, intro]: "finite A \ x \# mset_set A \ x \ A" blanchet@59813: by (induct A rule: finite_induct) simp_all blanchet@59813: haftmann@51548: context linorder haftmann@51548: begin haftmann@51548: haftmann@51548: definition sorted_list_of_multiset :: "'a multiset \ 'a list" haftmann@51548: where nipkow@59998: "sorted_list_of_multiset M = fold_mset insort [] M" haftmann@51548: haftmann@51548: lemma sorted_list_of_multiset_empty [simp]: haftmann@51548: "sorted_list_of_multiset {#} = []" haftmann@51548: by (simp add: sorted_list_of_multiset_def) haftmann@51548: haftmann@51548: lemma sorted_list_of_multiset_singleton [simp]: haftmann@51548: "sorted_list_of_multiset {#x#} = [x]" haftmann@51548: proof - haftmann@51548: interpret comp_fun_commute insort by (fact comp_fun_commute_insort) haftmann@51548: show ?thesis by (simp add: sorted_list_of_multiset_def) haftmann@51548: qed haftmann@51548: haftmann@51548: lemma sorted_list_of_multiset_insert [simp]: haftmann@51548: "sorted_list_of_multiset (M + {#x#}) = List.insort x (sorted_list_of_multiset M)" haftmann@51548: proof - haftmann@51548: interpret comp_fun_commute insort by (fact comp_fun_commute_insort) haftmann@51548: show ?thesis by (simp add: sorted_list_of_multiset_def) haftmann@51548: qed haftmann@51548: haftmann@51548: end haftmann@51548: nipkow@60515: lemma mset_sorted_list_of_multiset [simp]: nipkow@60515: "mset (sorted_list_of_multiset M) = M" nipkow@60513: by (induct M) simp_all haftmann@51548: nipkow@60515: lemma sorted_list_of_multiset_mset [simp]: nipkow@60515: "sorted_list_of_multiset (mset xs) = sort xs" nipkow@60513: by (induct xs) simp_all nipkow@60513: nipkow@60513: lemma finite_set_mset_mset_set[simp]: nipkow@60513: "finite A \ set_mset (mset_set A) = A" nipkow@60513: by (induct A rule: finite_induct) simp_all nipkow@60513: nipkow@60513: lemma infinite_set_mset_mset_set: nipkow@60513: "\ finite A \ set_mset (mset_set A) = {}" nipkow@60513: by simp haftmann@51548: haftmann@51548: lemma set_sorted_list_of_multiset [simp]: nipkow@60495: "set (sorted_list_of_multiset M) = set_mset M" nipkow@60513: by (induct M) (simp_all add: set_insort) nipkow@60513: nipkow@60513: lemma sorted_list_of_mset_set [simp]: nipkow@60513: "sorted_list_of_multiset (mset_set A) = sorted_list_of_set A" nipkow@60513: by (cases "finite A") (induct A rule: finite_induct, simp_all add: ac_simps) haftmann@51548: haftmann@51548: haftmann@60804: subsection \Replicate operation\ haftmann@60804: haftmann@60804: definition replicate_mset :: "nat \ 'a \ 'a multiset" where haftmann@60804: "replicate_mset n x = ((op + {#x#}) ^^ n) {#}" haftmann@60804: haftmann@60804: lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}" haftmann@60804: unfolding replicate_mset_def by simp haftmann@60804: haftmann@60804: lemma replicate_mset_Suc[simp]: "replicate_mset (Suc n) x = replicate_mset n x + {#x#}" haftmann@60804: unfolding replicate_mset_def by (induct n) (auto intro: add.commute) haftmann@60804: haftmann@60804: lemma in_replicate_mset[simp]: "x \# replicate_mset n y \ n > 0 \ x = y" haftmann@60804: unfolding replicate_mset_def by (induct n) simp_all haftmann@60804: haftmann@60804: lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)" haftmann@60804: unfolding replicate_mset_def by (induct n) simp_all haftmann@60804: haftmann@60804: lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})" haftmann@60804: by (auto split: if_splits) haftmann@60804: haftmann@60804: lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" haftmann@60804: by (induct n, simp_all) haftmann@60804: haftmann@60804: lemma count_le_replicate_mset_le: "n \ count M x \ replicate_mset n x \# M" haftmann@60804: by (auto simp add: assms mset_less_eqI) (metis count_replicate_mset subseteq_mset_def) haftmann@60804: haftmann@60804: lemma filter_eq_replicate_mset: "{#y \# D. y = x#} = replicate_mset (count D x) x" haftmann@60804: by (induct D) simp_all haftmann@60804: haftmann@61031: lemma replicate_count_mset_eq_filter_eq: haftmann@61031: "replicate (count (mset xs) k) k = filter (HOL.eq k) xs" haftmann@61031: by (induct xs) auto haftmann@61031: haftmann@60804: wenzelm@60500: subsection \Big operators\ haftmann@51548: haftmann@51548: no_notation times (infixl "*" 70) haftmann@51548: no_notation Groups.one ("1") haftmann@51548: haftmann@51548: locale comm_monoid_mset = comm_monoid haftmann@51548: begin haftmann@51548: haftmann@51548: definition F :: "'a multiset \ 'a" wenzelm@60606: where eq_fold: "F M = fold_mset f 1 M" wenzelm@60606: wenzelm@60606: lemma empty [simp]: "F {#} = 1" haftmann@51548: by (simp add: eq_fold) haftmann@51548: wenzelm@60678: lemma singleton [simp]: "F {#x#} = x" haftmann@51548: proof - haftmann@51548: interpret comp_fun_commute wenzelm@60678: by standard (simp add: fun_eq_iff left_commute) haftmann@51548: show ?thesis by (simp add: eq_fold) haftmann@51548: qed haftmann@51548: wenzelm@60606: lemma union [simp]: "F (M + N) = F M * F N" haftmann@51548: proof - haftmann@51548: interpret comp_fun_commute f wenzelm@60678: by standard (simp add: fun_eq_iff left_commute) wenzelm@60678: show ?thesis wenzelm@60678: by (induct N) (simp_all add: left_commute eq_fold) haftmann@51548: qed haftmann@51548: haftmann@51548: end haftmann@51548: wenzelm@61076: lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute (op + :: 'a multiset \ _ \ _)" wenzelm@60678: by standard (simp add: add_ac comp_def) blanchet@59813: blanchet@59813: declare comp_fun_commute.fold_mset_insert[OF comp_fun_commute_plus_mset, simp] blanchet@59813: nipkow@59998: lemma in_mset_fold_plus_iff[iff]: "x \# fold_mset (op +) M NN \ x \# M \ (\N. N \# NN \ x \# N)" blanchet@59813: by (induct NN) auto blanchet@59813: haftmann@51548: notation times (infixl "*" 70) haftmann@51548: notation Groups.one ("1") haftmann@51548: haftmann@54868: context comm_monoid_add haftmann@54868: begin haftmann@54868: haftmann@54868: definition msetsum :: "'a multiset \ 'a" wenzelm@60606: where "msetsum = comm_monoid_mset.F plus 0" haftmann@51548: haftmann@54868: sublocale msetsum!: comm_monoid_mset plus 0 ballarin@61566: rewrites "comm_monoid_mset.F plus 0 = msetsum" haftmann@51548: proof - haftmann@51548: show "comm_monoid_mset plus 0" .. haftmann@51548: from msetsum_def show "comm_monoid_mset.F plus 0 = msetsum" .. haftmann@51548: qed haftmann@51548: haftmann@60804: lemma (in semiring_1) msetsum_replicate_mset [simp]: haftmann@60804: "msetsum (replicate_mset n a) = of_nat n * a" haftmann@60804: by (induct n) (simp_all add: algebra_simps) haftmann@60804: haftmann@51548: lemma setsum_unfold_msetsum: nipkow@60513: "setsum f A = msetsum (image_mset f (mset_set A))" haftmann@51548: by (cases "finite A") (induct A rule: finite_induct, simp_all) haftmann@51548: haftmann@51548: end haftmann@51548: blanchet@59813: lemma msetsum_diff: wenzelm@61076: fixes M N :: "('a :: ordered_cancel_comm_monoid_diff) multiset" Mathias@60397: shows "N \# M \ msetsum (M - N) = msetsum M - msetsum N" Mathias@60397: by (metis add_diff_cancel_right' msetsum.union subset_mset.diff_add) blanchet@59813: nipkow@59949: lemma size_eq_msetsum: "size M = msetsum (image_mset (\_. 1) M)" nipkow@59949: proof (induct M) nipkow@59949: case empty then show ?case by simp nipkow@59949: next nipkow@59949: case (add M x) then show ?case nipkow@60495: by (cases "x \ set_mset M") nipkow@60495: (simp_all del: mem_set_mset_iff add: size_multiset_overloaded_eq setsum.distrib setsum.delta' insert_absorb, simp) nipkow@59949: qed nipkow@59949: nipkow@59949: blanchet@59813: abbreviation Union_mset :: "'a multiset multiset \ 'a multiset" where blanchet@59813: "Union_mset MM \ msetsum MM" blanchet@59813: blanchet@59813: notation (xsymbols) Union_mset ("\#_" [900] 900) blanchet@59813: nipkow@60495: lemma set_mset_Union_mset[simp]: "set_mset (\# MM) = (\M \ set_mset MM. set_mset M)" blanchet@59813: by (induct MM) auto blanchet@59813: blanchet@59813: lemma in_Union_mset_iff[iff]: "x \# \# MM \ (\M. M \# MM \ x \# M)" blanchet@59813: by (induct MM) auto blanchet@59813: haftmann@51548: syntax blanchet@58425: "_msetsum_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" haftmann@51548: ("(3SUM _:#_. _)" [0, 51, 10] 10) haftmann@51548: syntax (xsymbols) blanchet@58425: "_msetsum_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_add" haftmann@57518: ("(3\_\#_. _)" [0, 51, 10] 10) haftmann@51548: translations haftmann@57518: "SUM i :# A. b" == "CONST msetsum (CONST image_mset (\i. b) A)" haftmann@51548: haftmann@54868: context comm_monoid_mult haftmann@54868: begin haftmann@54868: haftmann@54868: definition msetprod :: "'a multiset \ 'a" wenzelm@60606: where "msetprod = comm_monoid_mset.F times 1" haftmann@54868: haftmann@54868: sublocale msetprod!: comm_monoid_mset times 1 ballarin@61566: rewrites "comm_monoid_mset.F times 1 = msetprod" haftmann@51548: proof - haftmann@51548: show "comm_monoid_mset times 1" .. wenzelm@60606: show "comm_monoid_mset.F times 1 = msetprod" using msetprod_def .. haftmann@51548: qed haftmann@51548: haftmann@51548: lemma msetprod_empty: haftmann@51548: "msetprod {#} = 1" haftmann@51548: by (fact msetprod.empty) haftmann@51548: haftmann@51548: lemma msetprod_singleton: haftmann@51548: "msetprod {#x#} = x" haftmann@51548: by (fact msetprod.singleton) haftmann@51548: haftmann@51548: lemma msetprod_Un: blanchet@58425: "msetprod (A + B) = msetprod A * msetprod B" haftmann@51548: by (fact msetprod.union) haftmann@51548: haftmann@60804: lemma msetprod_replicate_mset [simp]: haftmann@60804: "msetprod (replicate_mset n a) = a ^ n" haftmann@60804: by (induct n) (simp_all add: ac_simps) haftmann@60804: haftmann@51548: lemma setprod_unfold_msetprod: nipkow@60513: "setprod f A = msetprod (image_mset f (mset_set A))" haftmann@51548: by (cases "finite A") (induct A rule: finite_induct, simp_all) haftmann@51548: haftmann@51548: lemma msetprod_multiplicity: nipkow@60495: "msetprod M = setprod (\x. x ^ count M x) (set_mset M)" nipkow@59998: by (simp add: fold_mset_def setprod.eq_fold msetprod.eq_fold funpow_times_power comp_def) haftmann@51548: haftmann@51548: end haftmann@51548: haftmann@51548: syntax blanchet@58425: "_msetprod_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" haftmann@51548: ("(3PROD _:#_. _)" [0, 51, 10] 10) haftmann@51548: syntax (xsymbols) blanchet@58425: "_msetprod_image" :: "pttrn \ 'b set \ 'a \ 'a::comm_monoid_mult" haftmann@51548: ("(3\_\#_. _)" [0, 51, 10] 10) haftmann@51548: translations haftmann@57518: "PROD i :# A. b" == "CONST msetprod (CONST image_mset (\i. b) A)" haftmann@51548: haftmann@51548: lemma (in comm_semiring_1) dvd_msetprod: haftmann@51548: assumes "x \# A" haftmann@51548: shows "x dvd msetprod A" haftmann@51548: proof - haftmann@51548: from assms have "A = (A - {#x#}) + {#x#}" by simp haftmann@51548: then obtain B where "A = B + {#x#}" .. haftmann@51548: then show ?thesis by simp haftmann@51548: qed haftmann@51548: haftmann@51548: wenzelm@60500: subsection \Alternative representations\ wenzelm@60500: wenzelm@60500: subsubsection \Lists\ haftmann@51548: haftmann@39533: context linorder haftmann@39533: begin haftmann@39533: nipkow@60515: lemma mset_insort [simp]: nipkow@60515: "mset (insort_key k x xs) = {#x#} + mset xs" haftmann@37107: by (induct xs) (simp_all add: ac_simps) blanchet@58425: nipkow@60515: lemma mset_sort [simp]: nipkow@60515: "mset (sort_key k xs) = mset xs" haftmann@37107: by (induct xs) (simp_all add: ac_simps) haftmann@37107: wenzelm@60500: text \ haftmann@34943: This lemma shows which properties suffice to show that a function wenzelm@61585: \f\ with \f xs = ys\ behaves like sort. wenzelm@60500: \ haftmann@37074: haftmann@39533: lemma properties_for_sort_key: nipkow@60515: assumes "mset ys = mset xs" wenzelm@60606: and "\k. k \ set ys \ filter (\x. f k = f x) ys = filter (\x. f k = f x) xs" wenzelm@60606: and "sorted (map f ys)" haftmann@39533: shows "sort_key f xs = ys" wenzelm@60606: using assms wenzelm@46921: proof (induct xs arbitrary: ys) haftmann@34943: case Nil then show ?case by simp haftmann@34943: next haftmann@34943: case (Cons x xs) haftmann@39533: from Cons.prems(2) have haftmann@40305: "\k \ set ys. filter (\x. f k = f x) (remove1 x ys) = filter (\x. f k = f x) xs" haftmann@39533: by (simp add: filter_remove1) haftmann@39533: with Cons.prems have "sort_key f xs = remove1 x ys" haftmann@39533: by (auto intro!: Cons.hyps simp add: sorted_map_remove1) haftmann@39533: moreover from Cons.prems have "x \ set ys" haftmann@39533: by (auto simp add: mem_set_multiset_eq intro!: ccontr) haftmann@39533: ultimately show ?case using Cons.prems by (simp add: insort_key_remove1) haftmann@34943: qed haftmann@34943: haftmann@39533: lemma properties_for_sort: nipkow@60515: assumes multiset: "mset ys = mset xs" wenzelm@60606: and "sorted ys" haftmann@39533: shows "sort xs = ys" haftmann@39533: proof (rule properties_for_sort_key) nipkow@60515: from multiset show "mset ys = mset xs" . wenzelm@60500: from \sorted ys\ show "sorted (map (\x. x) ys)" by simp wenzelm@60678: from multiset have "length (filter (\y. k = y) ys) = length (filter (\x. k = x) xs)" for k nipkow@60515: by (rule mset_eq_length_filter) wenzelm@60678: then have "replicate (length (filter (\y. k = y) ys)) k = wenzelm@60678: replicate (length (filter (\x. k = x) xs)) k" for k haftmann@39533: by simp wenzelm@60678: then show "k \ set ys \ filter (\y. k = y) ys = filter (\x. k = x) xs" for k haftmann@39533: by (simp add: replicate_length_filter) haftmann@39533: qed haftmann@39533: haftmann@61031: lemma sort_key_inj_key_eq: haftmann@61031: assumes mset_equal: "mset xs = mset ys" haftmann@61031: and "inj_on f (set xs)" haftmann@61031: and "sorted (map f ys)" haftmann@61031: shows "sort_key f xs = ys" haftmann@61031: proof (rule properties_for_sort_key) haftmann@61031: from mset_equal haftmann@61031: show "mset ys = mset xs" by simp wenzelm@61188: from \sorted (map f ys)\ haftmann@61031: show "sorted (map f ys)" . haftmann@61031: show "[x\ys . f k = f x] = [x\xs . f k = f x]" if "k \ set ys" for k haftmann@61031: proof - haftmann@61031: from mset_equal haftmann@61031: have set_equal: "set xs = set ys" by (rule mset_eq_setD) haftmann@61031: with that have "insert k (set ys) = set ys" by auto wenzelm@61188: with \inj_on f (set xs)\ have inj: "inj_on f (insert k (set ys))" haftmann@61031: by (simp add: set_equal) haftmann@61031: from inj have "[x\ys . f k = f x] = filter (HOL.eq k) ys" haftmann@61031: by (auto intro!: inj_on_filter_key_eq) haftmann@61031: also have "\ = replicate (count (mset ys) k) k" haftmann@61031: by (simp add: replicate_count_mset_eq_filter_eq) haftmann@61031: also have "\ = replicate (count (mset xs) k) k" haftmann@61031: using mset_equal by simp haftmann@61031: also have "\ = filter (HOL.eq k) xs" haftmann@61031: by (simp add: replicate_count_mset_eq_filter_eq) haftmann@61031: also have "\ = [x\xs . f k = f x]" haftmann@61031: using inj by (auto intro!: inj_on_filter_key_eq [symmetric] simp add: set_equal) haftmann@61031: finally show ?thesis . haftmann@61031: qed haftmann@61031: qed haftmann@61031: haftmann@61031: lemma sort_key_eq_sort_key: haftmann@61031: assumes "mset xs = mset ys" haftmann@61031: and "inj_on f (set xs)" haftmann@61031: shows "sort_key f xs = sort_key f ys" haftmann@61031: by (rule sort_key_inj_key_eq) (simp_all add: assms) haftmann@61031: haftmann@40303: lemma sort_key_by_quicksort: haftmann@40303: "sort_key f xs = sort_key f [x\xs. f x < f (xs ! (length xs div 2))] haftmann@40303: @ [x\xs. f x = f (xs ! (length xs div 2))] haftmann@40303: @ sort_key f [x\xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs") haftmann@40303: proof (rule properties_for_sort_key) nipkow@60515: show "mset ?rhs = mset ?lhs" nipkow@60515: by (rule multiset_eqI) (auto simp add: mset_filter) haftmann@40303: show "sorted (map f ?rhs)" haftmann@40303: by (auto simp add: sorted_append intro: sorted_map_same) haftmann@40303: next haftmann@40305: fix l haftmann@40305: assume "l \ set ?rhs" haftmann@40346: let ?pivot = "f (xs ! (length xs div 2))" haftmann@40346: have *: "\x. f l = f x \ f x = f l" by auto haftmann@40306: have "[x \ sort_key f xs . f x = f l] = [x \ xs. f x = f l]" haftmann@40305: unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) haftmann@40346: with * have **: "[x \ sort_key f xs . f l = f x] = [x \ xs. f l = f x]" by simp haftmann@40346: have "\x P. P (f x) ?pivot \ f l = f x \ P (f l) ?pivot \ f l = f x" by auto haftmann@40346: then have "\P. [x \ sort_key f xs . P (f x) ?pivot \ f l = f x] = haftmann@40346: [x \ sort_key f xs. P (f l) ?pivot \ f l = f x]" by simp haftmann@40346: note *** = this [of "op <"] this [of "op >"] this [of "op ="] haftmann@40306: show "[x \ ?rhs. f l = f x] = [x \ ?lhs. f l = f x]" haftmann@40305: proof (cases "f l" ?pivot rule: linorder_cases) wenzelm@46730: case less wenzelm@46730: then have "f l \ ?pivot" and "\ f l > ?pivot" by auto wenzelm@46730: with less show ?thesis haftmann@40346: by (simp add: filter_sort [symmetric] ** ***) haftmann@40305: next haftmann@40306: case equal then show ?thesis haftmann@40346: by (simp add: * less_le) haftmann@40305: next wenzelm@46730: case greater wenzelm@46730: then have "f l \ ?pivot" and "\ f l < ?pivot" by auto wenzelm@46730: with greater show ?thesis haftmann@40346: by (simp add: filter_sort [symmetric] ** ***) haftmann@40306: qed haftmann@40303: qed haftmann@40303: haftmann@40303: lemma sort_by_quicksort: haftmann@40303: "sort xs = sort [x\xs. x < xs ! (length xs div 2)] haftmann@40303: @ [x\xs. x = xs ! (length xs div 2)] haftmann@40303: @ sort [x\xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") haftmann@40303: using sort_key_by_quicksort [of "\x. x", symmetric] by simp haftmann@40303: wenzelm@60500: text \A stable parametrized quicksort\ haftmann@40347: haftmann@40347: definition part :: "('b \ 'a) \ 'a \ 'b list \ 'b list \ 'b list \ 'b list" where haftmann@40347: "part f pivot xs = ([x \ xs. f x < pivot], [x \ xs. f x = pivot], [x \ xs. pivot < f x])" haftmann@40347: haftmann@40347: lemma part_code [code]: haftmann@40347: "part f pivot [] = ([], [], [])" haftmann@40347: "part f pivot (x # xs) = (let (lts, eqs, gts) = part f pivot xs; x' = f x in haftmann@40347: if x' < pivot then (x # lts, eqs, gts) haftmann@40347: else if x' > pivot then (lts, eqs, x # gts) haftmann@40347: else (lts, x # eqs, gts))" haftmann@40347: by (auto simp add: part_def Let_def split_def) haftmann@40347: haftmann@40347: lemma sort_key_by_quicksort_code [code]: wenzelm@60606: "sort_key f xs = wenzelm@60606: (case xs of wenzelm@60606: [] \ [] haftmann@40347: | [x] \ xs haftmann@40347: | [x, y] \ (if f x \ f y then xs else [y, x]) wenzelm@60606: | _ \ wenzelm@60606: let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs wenzelm@60606: in sort_key f lts @ eqs @ sort_key f gts)" haftmann@40347: proof (cases xs) haftmann@40347: case Nil then show ?thesis by simp haftmann@40347: next wenzelm@46921: case (Cons _ ys) note hyps = Cons show ?thesis wenzelm@46921: proof (cases ys) haftmann@40347: case Nil with hyps show ?thesis by simp haftmann@40347: next wenzelm@46921: case (Cons _ zs) note hyps = hyps Cons show ?thesis wenzelm@46921: proof (cases zs) haftmann@40347: case Nil with hyps show ?thesis by auto haftmann@40347: next blanchet@58425: case Cons haftmann@40347: from sort_key_by_quicksort [of f xs] haftmann@40347: have "sort_key f xs = (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs haftmann@40347: in sort_key f lts @ eqs @ sort_key f gts)" haftmann@40347: by (simp only: split_def Let_def part_def fst_conv snd_conv) haftmann@40347: with hyps Cons show ?thesis by (simp only: list.cases) haftmann@40347: qed haftmann@40347: qed haftmann@40347: qed haftmann@40347: haftmann@39533: end haftmann@39533: haftmann@40347: hide_const (open) part haftmann@40347: nipkow@60515: lemma mset_remdups_le: "mset (remdups xs) \# mset xs" Mathias@60397: by (induct xs) (auto intro: subset_mset.order_trans) haftmann@34943: nipkow@60515: lemma mset_update: nipkow@60515: "i < length ls \ mset (ls[i := v]) = mset ls - {#ls ! i#} + {#v#}" haftmann@34943: proof (induct ls arbitrary: i) haftmann@34943: case Nil then show ?case by simp haftmann@34943: next haftmann@34943: case (Cons x xs) haftmann@34943: show ?case haftmann@34943: proof (cases i) haftmann@34943: case 0 then show ?thesis by simp haftmann@34943: next haftmann@34943: case (Suc i') haftmann@34943: with Cons show ?thesis haftmann@34943: apply simp haftmann@57512: apply (subst add.assoc) haftmann@57512: apply (subst add.commute [of "{#v#}" "{#x#}"]) haftmann@57512: apply (subst add.assoc [symmetric]) haftmann@34943: apply simp haftmann@34943: apply (rule mset_le_multiset_union_diff_commute) nipkow@60515: apply (simp add: mset_le_single nth_mem_mset) haftmann@34943: done haftmann@34943: qed haftmann@34943: qed haftmann@34943: nipkow@60515: lemma mset_swap: haftmann@34943: "i < length ls \ j < length ls \ nipkow@60515: mset (ls[j := ls ! i, i := ls ! j]) = mset ls" nipkow@60515: by (cases "i = j") (simp_all add: mset_update nth_mem_mset) haftmann@34943: haftmann@34943: wenzelm@60500: subsection \The multiset order\ wenzelm@60500: wenzelm@60500: subsubsection \Well-foundedness\ wenzelm@10249: wenzelm@60606: definition mult1 :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where haftmann@37765: "mult1 r = {(N, M). \a M0 K. M = M0 + {#a#} \ N = M0 + K \ wenzelm@60607: (\b. b \# K \ (b, a) \ r)}" wenzelm@60606: wenzelm@60606: definition mult :: "('a \ 'a) set \ ('a multiset \ 'a multiset) set" where haftmann@37765: "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: wenzelm@60608: lemma less_add: wenzelm@60608: assumes mult1: "(N, M0 + {#a#}) \ mult1 r" wenzelm@60608: shows wenzelm@60608: "(\M. (M, M0) \ mult1 r \ N = M + {#a#}) \ wenzelm@60608: (\K. (\b. b \# K \ (b, a) \ r) \ N = M0 + K)" wenzelm@60608: proof - wenzelm@60607: 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" wenzelm@60608: obtain a' M0' K where M0: "M0 + {#a#} = M0' + {#a'#}" wenzelm@60608: and N: "N = M0' + K" wenzelm@60608: and r: "?r K a'" wenzelm@60608: using mult1 unfolding mult1_def by auto wenzelm@60608: show ?thesis (is "?case1 \ ?case2") wenzelm@60606: proof - wenzelm@60606: from M0 consider "M0 = M0'" "a = a'" wenzelm@60606: | K' where "M0 = K' + {#a'#}" "M0' = K' + {#a#}" wenzelm@60606: by atomize_elim (simp only: add_eq_conv_ex) wenzelm@18258: then show ?thesis wenzelm@60606: proof cases wenzelm@60606: case 1 nipkow@11464: with N r have "?r K a \ N = M0 + K" by simp wenzelm@60606: then have ?case2 .. wenzelm@60606: then show ?thesis .. wenzelm@10249: next wenzelm@60606: case 2 wenzelm@60606: from N 2(2) have n: "N = K' + K + {#a#}" by (simp add: ac_simps) wenzelm@60606: with r 2(1) have "?R (K' + K) M0" by blast wenzelm@60608: with n have ?case1 by (simp add: mult1_def) wenzelm@60606: then show ?thesis .. wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: wenzelm@60608: lemma all_accessible: wenzelm@60608: assumes "wf r" wenzelm@60608: shows "\M. M \ Wellfounded.acc (mult1 r)" wenzelm@10249: proof wenzelm@10249: let ?R = "mult1 r" haftmann@54295: let ?W = "Wellfounded.acc ?R" wenzelm@10249: { wenzelm@10249: fix M M0 a berghofe@23751: assume M0: "M0 \ ?W" wenzelm@60606: and wf_hyp: "\b. (b, a) \ r \ (\M \ ?W. M + {#b#} \ ?W)" wenzelm@60606: 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" wenzelm@60608: then consider M where "(M, M0) \ ?R" "N = M + {#a#}" wenzelm@60608: | K where "\b. b \# K \ (b, a) \ r" "N = M0 + K" wenzelm@60608: by atomize_elim (rule less_add) berghofe@23751: then show "N \ ?W" wenzelm@60608: proof cases wenzelm@60608: case 1 wenzelm@60606: from acc_hyp have "(M, M0) \ ?R \ M + {#a#} \ ?W" .. wenzelm@60500: from this and \(M, M0) \ ?R\ have "M + {#a#} \ ?W" .. wenzelm@60608: then show "N \ ?W" by (simp only: \N = M + {#a#}\) wenzelm@10249: next wenzelm@60608: case 2 wenzelm@60608: from this(1) 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" .. haftmann@57512: then show "M0 + (K + {#x#}) \ ?W" by (simp only: add.assoc) wenzelm@10249: qed wenzelm@60608: then show "N \ ?W" by (simp only: 2(2)) wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: } note tedious_reasoning = this wenzelm@10249: wenzelm@60608: show "M \ ?W" for M 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" wenzelm@60608: from \wf r\ have "\M \ ?W. M + {#a#} \ ?W" wenzelm@10249: proof induct wenzelm@10249: fix a wenzelm@60606: 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 wenzelm@60500: from this and \M \ ?W\ show "M + {#a#} \ ?W" .. wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: wenzelm@60606: theorem wf_mult1: "wf r \ wf (mult1 r)" nipkow@26178: by (rule acc_wfI) (rule all_accessible) wenzelm@10249: wenzelm@60606: 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@60500: subsubsection \Closure-free presentation\ wenzelm@60500: wenzelm@60500: text \One direction.\ wenzelm@10249: wenzelm@10249: lemma mult_implies_one_step: wenzelm@60606: "trans r \ (M, N) \ mult r \ nipkow@11464: \I J K. N = I + J \ M = I + K \ J \ {#} \ nipkow@60495: (\k \ set_mset K. \j \ set_mset J. (k, j) \ r)" nipkow@60495: apply (unfold mult_def mult1_def set_mset_def) nipkow@26178: apply (erule converse_trancl_induct, clarify) nipkow@26178: apply (rule_tac x = M0 in exI, simp, clarify) wenzelm@60607: 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) haftmann@57512: apply (simp (no_asm_simp) add: add.assoc [symmetric]) wenzelm@59807: apply (drule_tac f = "\M. M - {#a#}" and x="S + T" for S T 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 wenzelm@60607: 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@39302: apply (simp add: multiset_eq_iff split: nat_diff_split) nipkow@26178: apply (rule conjI) wenzelm@59807: apply (drule_tac f = "\M. M - {#a#}" and x="S + T" for S T in arg_cong, simp) nipkow@39302: apply (simp add: multiset_eq_iff split: nat_diff_split) nipkow@26178: apply (simp (no_asm_use) add: trans_def) nipkow@26178: apply blast wenzelm@60607: apply (subgoal_tac "a \# (M0 + {#a#})") nipkow@26178: apply simp nipkow@26178: apply (simp (no_asm)) nipkow@26178: done wenzelm@10249: wenzelm@10249: lemma one_step_implies_mult_aux: wenzelm@60678: "\I J K. size J = n \ J \ {#} \ (\k \ set_mset K. \j \ set_mset J. (k, j) \ r) wenzelm@60678: \ (I + K, I + J) \ mult r" wenzelm@60678: apply (induct n) wenzelm@60678: apply 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@60495: apply (simp add: mult1_def set_mset_def, blast) wenzelm@60500: txt \Now we know @{term "J' \ {#}"}.\ nipkow@26178: apply (cut_tac M = K and P = "\x. (x, a) \ r" in multiset_partition) nipkow@60495: apply (erule_tac P = "\k \ set_mset K. P k" for P in rev_mp) nipkow@26178: apply (erule ssubst) nipkow@26178: apply (simp add: Ball_def, auto) nipkow@26178: apply (subgoal_tac wenzelm@60607: "((I + {# x \# K. (x, a) \ r #}) + {# x \# K. (x, a) \ r #}, wenzelm@60607: (I + {# x \# K. (x, a) \ r #}) + J') \ mult r") nipkow@26178: prefer 2 nipkow@26178: apply force haftmann@57512: apply (simp (no_asm_use) add: add.assoc [symmetric] mult_def) nipkow@26178: apply (erule trancl_trans) nipkow@26178: apply (rule r_into_trancl) nipkow@60495: apply (simp add: mult1_def set_mset_def) nipkow@26178: apply (rule_tac x = a in exI) nipkow@26178: apply (rule_tac x = "I + J'" in exI) haftmann@57514: apply (simp add: ac_simps) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma one_step_implies_mult: wenzelm@60606: "trans r \ J \ {#} \ \k \ set_mset K. \j \ set_mset J. (k, j) \ r wenzelm@60606: \ (I + K, I + J) \ mult r" nipkow@26178: using one_step_implies_mult_aux by blast wenzelm@10249: wenzelm@10249: wenzelm@60500: subsubsection \Partial-order properties\ wenzelm@10249: wenzelm@61076: definition less_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#<#" 50) where blanchet@59958: "M' #<# M \ (M', M) \ mult {(x', x). x' < x}" blanchet@59958: wenzelm@61076: definition le_multiset :: "'a::order multiset \ 'a multiset \ bool" (infix "#<=#" 50) where blanchet@59958: "M' #<=# M \ M' #<# M \ M' = M" blanchet@59958: blanchet@59958: notation (xsymbols) less_multiset (infix "#\#" 50) blanchet@59958: notation (xsymbols) le_multiset (infix "#\#" 50) wenzelm@10249: haftmann@35268: interpretation multiset_order: order le_multiset less_multiset haftmann@35268: proof - wenzelm@60606: have irrefl: "\ M #\# M" for M :: "'a multiset" haftmann@35268: proof blanchet@59958: assume "M #\# M" haftmann@35268: then have MM: "(M, M) \ mult {(x, y). x < y}" by (simp add: less_multiset_def) haftmann@35268: have "trans {(x'::'a, x). x' < x}" haftmann@35268: by (rule transI) simp haftmann@35268: moreover note MM haftmann@35268: ultimately have "\I J K. M = I + J \ M = I + K nipkow@60495: \ J \ {#} \ (\k\set_mset K. \j\set_mset J. (k, j) \ {(x, y). x < y})" haftmann@35268: by (rule mult_implies_one_step) haftmann@35268: then obtain I J K where "M = I + J" and "M = I + K" nipkow@60495: and "J \ {#}" and "(\k\set_mset K. \j\set_mset J. (k, j) \ {(x, y). x < y})" by blast wenzelm@60678: then have *: "K \ {#}" and **: "\k\set_mset K. \j\set_mset K. k < j" by auto nipkow@60495: have "finite (set_mset K)" by simp wenzelm@60678: moreover note ** nipkow@60495: ultimately have "set_mset K = {}" haftmann@35268: by (induct rule: finite_induct) (auto intro: order_less_trans) wenzelm@60678: with * show False by simp haftmann@35268: qed wenzelm@60678: have trans: "K #\# M \ M #\# N \ K #\# N" for K M N :: "'a multiset" haftmann@35268: unfolding less_multiset_def mult_def by (blast intro: trancl_trans) wenzelm@46921: show "class.order (le_multiset :: 'a multiset \ _) less_multiset" wenzelm@60678: by standard (auto simp add: le_multiset_def irrefl dest: trans) haftmann@35268: qed wenzelm@10249: wenzelm@60678: lemma mult_less_irrefl [elim!]: wenzelm@60678: fixes M :: "'a::order multiset" wenzelm@60678: shows "M #\# M \ R" wenzelm@46730: by simp haftmann@26567: wenzelm@10249: wenzelm@60500: subsubsection \Monotonicity of multiset union\ wenzelm@10249: wenzelm@60606: lemma mult1_union: "(B, D) \ mult1 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) haftmann@57512: apply (simp add: add.assoc) nipkow@26178: done wenzelm@10249: wenzelm@60606: 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) noschinl@40249: apply (blast intro: mult1_union) noschinl@40249: apply (blast intro: mult1_union trancl_trans) nipkow@26178: done wenzelm@10249: wenzelm@60606: lemma union_less_mono1: "B #\# D \ B + C #\# D + (C::'a::order multiset)" haftmann@57512: apply (subst add.commute [of B C]) haftmann@57512: apply (subst add.commute [of D C]) nipkow@26178: apply (erule union_less_mono2) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma union_less_mono: wenzelm@60606: fixes A B C D :: "'a::order multiset" wenzelm@60606: shows "A #\# C \ B #\# D \ A + B #\# C + D" haftmann@35268: by (blast intro!: union_less_mono1 union_less_mono2 multiset_order.less_trans) wenzelm@10249: haftmann@35268: interpretation multiset_order: ordered_ab_semigroup_add plus le_multiset less_multiset wenzelm@60678: by standard (auto simp add: le_multiset_def intro: union_less_mono2) wenzelm@26145: paulson@15072: wenzelm@60500: subsubsection \Termination proofs with multiset orders\ krauss@29125: krauss@29125: lemma multi_member_skip: "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" haftmann@37765: 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: nipkow@60495: "(set_mset A, set_mset 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: nipkow@60495: "(set_mset A, set_mset 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" nipkow@60495: shows "\A B Z. X = A + Z \ Y = B + Z \ ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" krauss@29125: using assms wenzelm@60606: 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 blanchet@58425: [simp]: "X = A + Z" "Y = B + Z" nipkow@60495: and 1[simp]: "(set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#})" krauss@29125: by auto wenzelm@60606: from pw_leq_step consider "x = y" | "(x, y) \ pair_less" krauss@29125: unfolding pair_leq_def by auto krauss@29125: thus ?case wenzelm@60606: proof cases wenzelm@60606: case [simp]: 1 wenzelm@60606: have "{#x#} + X = A + ({#y#}+Z) \ {#y#} + Y = B + ({#y#}+Z) \ wenzelm@60606: ((set_mset A, set_mset B) \ max_strict \ (B = {#} \ A = {#}))" haftmann@57514: by (auto simp: ac_simps) wenzelm@60606: thus ?thesis by blast krauss@29125: next wenzelm@60606: case 2 krauss@29125: let ?A' = "{#x#} + A" and ?B' = "{#y#} + B" krauss@29125: have "{#x#} + X = ?A' + Z" krauss@29125: "{#y#} + Y = ?B' + Z" haftmann@57514: by (auto simp add: ac_simps) blanchet@58425: moreover have nipkow@60495: "(set_mset ?A', set_mset ?B') \ max_strict" wenzelm@60606: using 1 2 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: blanchet@58425: lemma krauss@29125: assumes pwleq: "pw_leq Z Z'" nipkow@60495: shows ms_strictI: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_strict" wenzelm@60606: and ms_weakI1: "(set_mset A, set_mset B) \ max_strict \ (Z + A, Z' + B) \ ms_weak" wenzelm@60606: and ms_weakI2: "(Z + {#}, Z' + {#}) \ ms_weak" krauss@29125: proof - blanchet@58425: from pw_leq_split[OF pwleq] krauss@29125: obtain A' B' Z'' krauss@29125: where [simp]: "Z = A' + Z''" "Z' = B' + Z''" nipkow@60495: and mx_or_empty: "(set_mset A', set_mset B') \ max_strict \ (A' = {#} \ B' = {#})" krauss@29125: by blast krauss@29125: { nipkow@60495: assume max: "(set_mset A, set_mset B) \ max_strict" krauss@29125: from mx_or_empty krauss@29125: have "(Z'' + (A + A'), Z'' + (B + B')) \ ms_strict" krauss@29125: proof nipkow@60495: assume max': "(set_mset A', set_mset B') \ max_strict" nipkow@60495: with max have "(set_mset (A + A'), set_mset (B + B')) \ max_strict" krauss@29125: by (auto simp: max_strict_def intro: max_ext_additive) blanchet@58425: 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 wenzelm@60606: thus "(Z + A, Z' + B) \ ms_strict" by (simp add: ac_simps) 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) haftmann@57514: thus "(Z + {#}, Z' + {#}) \ ms_weak" by (simp add:ac_simps) krauss@29125: qed krauss@29125: nipkow@39301: lemma empty_neutral: "{#} + x = x" "x + {#} = x" krauss@29125: and nonempty_plus: "{# x #} + rs \ {#}" krauss@29125: and nonempty_single: "{# x #} \ {#}" krauss@29125: by auto krauss@29125: wenzelm@60500: setup \ wenzelm@60606: let wenzelm@60606: fun msetT T = Type (@{type_name multiset}, [T]); wenzelm@60606: wenzelm@60606: fun mk_mset T [] = Const (@{const_abbrev Mempty}, msetT T) wenzelm@60606: | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x wenzelm@60606: | mk_mset T (x :: xs) = wenzelm@60606: Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $ wenzelm@60606: mk_mset T [x] $ mk_mset T xs wenzelm@60606: wenzelm@60752: fun mset_member_tac ctxt m i = wenzelm@60606: if m <= 0 then wenzelm@60752: resolve_tac ctxt @{thms multi_member_this} i ORELSE wenzelm@60752: resolve_tac ctxt @{thms multi_member_last} i wenzelm@60606: else wenzelm@60752: resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i wenzelm@60752: wenzelm@60752: fun mset_nonempty_tac ctxt = wenzelm@60752: resolve_tac ctxt @{thms nonempty_plus} ORELSE' wenzelm@60752: resolve_tac ctxt @{thms nonempty_single} krauss@29125: wenzelm@60606: fun regroup_munion_conv ctxt = wenzelm@60606: Function_Lib.regroup_conv ctxt @{const_abbrev Mempty} @{const_name plus} wenzelm@60606: (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral})) wenzelm@60606: wenzelm@60752: fun unfold_pwleq_tac ctxt i = wenzelm@60752: (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st)) wenzelm@60752: ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i) wenzelm@60752: ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i) wenzelm@60606: wenzelm@60606: val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union}, wenzelm@60606: @{thm Un_insert_left}, @{thm Un_empty_left}] wenzelm@60606: in wenzelm@60606: ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset wenzelm@60606: { wenzelm@60606: msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, wenzelm@60606: mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, wenzelm@60606: mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps, wenzelm@60606: smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1}, wenzelm@60752: reduction_pair = @{thm ms_reduction_pair} wenzelm@60606: }) wenzelm@60606: end wenzelm@60500: \ wenzelm@60500: wenzelm@60500: wenzelm@60500: subsection \Legacy theorem bindings\ haftmann@34943: nipkow@39302: lemmas multi_count_eq = multiset_eq_iff [symmetric] haftmann@34943: haftmann@34943: lemma union_commute: "M + N = N + (M::'a multiset)" haftmann@57512: by (fact add.commute) haftmann@34943: haftmann@34943: lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" haftmann@57512: by (fact add.assoc) haftmann@34943: haftmann@34943: lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" haftmann@57512: by (fact add.left_commute) haftmann@34943: haftmann@34943: lemmas union_ac = union_assoc union_commute union_lcomm haftmann@34943: haftmann@34943: lemma union_right_cancel: "M + K = N + K \ M = (N::'a multiset)" haftmann@34943: by (fact add_right_cancel) haftmann@34943: haftmann@34943: lemma union_left_cancel: "K + M = K + N \ M = (N::'a multiset)" haftmann@34943: by (fact add_left_cancel) haftmann@34943: haftmann@34943: lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \ X = Y" haftmann@59557: by (fact add_left_imp_eq) haftmann@34943: Mathias@60397: lemma mset_less_trans: "(M::'a multiset) <# K \ K <# N \ M <# N" Mathias@60397: by (fact subset_mset.less_trans) haftmann@35268: haftmann@35268: lemma multiset_inter_commute: "A #\ B = B #\ A" Mathias@60397: by (fact subset_mset.inf.commute) haftmann@35268: haftmann@35268: lemma multiset_inter_assoc: "A #\ (B #\ C) = A #\ B #\ C" Mathias@60397: by (fact subset_mset.inf.assoc [symmetric]) haftmann@35268: haftmann@35268: lemma multiset_inter_left_commute: "A #\ (B #\ C) = B #\ (A #\ C)" Mathias@60397: by (fact subset_mset.inf.left_commute) haftmann@35268: haftmann@35268: lemmas multiset_inter_ac = haftmann@35268: multiset_inter_commute haftmann@35268: multiset_inter_assoc haftmann@35268: multiset_inter_left_commute haftmann@35268: wenzelm@60606: lemma mult_less_not_refl: "\ M #\# (M::'a::order multiset)" haftmann@35268: by (fact multiset_order.less_irrefl) haftmann@35268: wenzelm@60606: lemma mult_less_trans: "K #\# M \ M #\# N \ K #\# (N::'a::order multiset)" haftmann@35268: by (fact multiset_order.less_trans) blanchet@58425: wenzelm@60606: lemma mult_less_not_sym: "M #\# N \ \ N #\# (M::'a::order multiset)" haftmann@35268: by (fact multiset_order.less_not_sym) haftmann@35268: wenzelm@60606: lemma mult_less_asym: "M #\# N \ (\ P \ N #\# (M::'a::order multiset)) \ P" haftmann@35268: by (fact multiset_order.less_asym) haftmann@34943: wenzelm@60500: declaration \ wenzelm@60606: let wenzelm@60606: fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') = wenzelm@60606: let wenzelm@60606: val (maybe_opt, ps) = wenzelm@60606: Nitpick_Model.dest_plain_fun t' wenzelm@60606: ||> op ~~ wenzelm@60606: ||> map (apsnd (snd o HOLogic.dest_number)) wenzelm@60606: fun elems_for t = wenzelm@60606: (case AList.lookup (op =) ps t of wenzelm@60606: SOME n => replicate n t wenzelm@60606: | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]) wenzelm@60606: in wenzelm@60606: (case maps elems_for (all_values elem_T) @ blanchet@61333: (if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of wenzelm@60606: [] => Const (@{const_name zero_class.zero}, T) wenzelm@60606: | ts => wenzelm@60606: foldl1 (fn (t1, t2) => wenzelm@60606: Const (@{const_name plus_class.plus}, T --> T --> T) $ t1 $ t2) wenzelm@60606: (map (curry (op $) (Const (@{const_name single}, elem_T --> T))) ts)) wenzelm@60606: end wenzelm@60606: | multiset_postproc _ _ _ _ t = t wenzelm@60606: in Nitpick_Model.register_term_postprocessor @{typ "'a multiset"} multiset_postproc end wenzelm@60500: \ wenzelm@60500: wenzelm@60500: wenzelm@60500: subsection \Naive implementation using lists\ haftmann@51600: nipkow@60515: code_datatype mset haftmann@51600: wenzelm@60606: lemma [code]: "{#} = mset []" haftmann@51600: by simp haftmann@51600: wenzelm@60606: lemma [code]: "{#x#} = mset [x]" haftmann@51600: by simp haftmann@51600: wenzelm@60606: lemma union_code [code]: "mset xs + mset ys = mset (xs @ ys)" haftmann@51600: by simp haftmann@51600: wenzelm@60606: lemma [code]: "image_mset f (mset xs) = mset (map f xs)" nipkow@60515: by (simp add: mset_map) haftmann@51600: wenzelm@60606: lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)" nipkow@60515: by (simp add: mset_filter) haftmann@51600: wenzelm@60606: lemma [code]: "mset xs - mset ys = mset (fold remove1 ys xs)" haftmann@51600: by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute) haftmann@51600: haftmann@51600: lemma [code]: nipkow@60515: "mset xs #\ mset ys = nipkow@60515: mset (snd (fold (\x (ys, zs). haftmann@51600: if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, [])))" haftmann@51600: proof - nipkow@60515: have "\zs. mset (snd (fold (\x (ys, zs). haftmann@51600: if x \ set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) = nipkow@60515: (mset xs #\ mset ys) + mset zs" haftmann@51623: by (induct xs arbitrary: ys) haftmann@51623: (auto simp add: mem_set_multiset_eq inter_add_right1 inter_add_right2 ac_simps) haftmann@51623: then show ?thesis by simp haftmann@51623: qed haftmann@51623: haftmann@51623: lemma [code]: nipkow@60515: "mset xs #\ mset ys = haftmann@61424: mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" haftmann@51623: proof - haftmann@61424: have "\zs. mset (case_prod append (fold (\x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) = nipkow@60515: (mset xs #\ mset ys) + mset zs" haftmann@51623: by (induct xs arbitrary: ys) (simp_all add: multiset_eq_iff) haftmann@51600: then show ?thesis by simp haftmann@51600: qed haftmann@51600: blanchet@59813: declare in_multiset_in_set [code_unfold] haftmann@51600: wenzelm@60606: lemma [code]: "count (mset xs) x = fold (\y. if x = y then Suc else id) xs 0" haftmann@51600: proof - nipkow@60515: have "\n. fold (\y. if x = y then Suc else id) xs n = count (mset xs) x + n" haftmann@51600: by (induct xs) simp_all haftmann@51600: then show ?thesis by simp haftmann@51600: qed haftmann@51600: nipkow@60515: declare set_mset_mset [code] nipkow@60515: nipkow@60515: declare sorted_list_of_multiset_mset [code] haftmann@51600: wenzelm@61585: lemma [code]: \ \not very efficient, but representation-ignorant!\ nipkow@60515: "mset_set A = mset (sorted_list_of_set A)" haftmann@51600: apply (cases "finite A") haftmann@51600: apply simp_all haftmann@51600: apply (induct A rule: finite_induct) blanchet@59813: apply (simp_all add: add.commute) haftmann@51600: done haftmann@51600: nipkow@60515: declare size_mset [code] haftmann@51600: blanchet@58425: fun ms_lesseq_impl :: "'a list \ 'a list \ bool option" where nipkow@55808: "ms_lesseq_impl [] ys = Some (ys \ [])" blanchet@58425: | "ms_lesseq_impl (Cons x xs) ys = (case List.extract (op = x) ys of nipkow@55808: None \ None nipkow@55808: | Some (ys1,_,ys2) \ ms_lesseq_impl xs (ys1 @ ys2))" nipkow@55808: nipkow@60515: lemma ms_lesseq_impl: "(ms_lesseq_impl xs ys = None \ \ mset xs \# mset ys) \ nipkow@60515: (ms_lesseq_impl xs ys = Some True \ mset xs <# mset ys) \ nipkow@60515: (ms_lesseq_impl xs ys = Some False \ mset xs = mset ys)" nipkow@55808: proof (induct xs arbitrary: ys) nipkow@55808: case (Nil ys) nipkow@55808: show ?case by (auto simp: mset_less_empty_nonempty) nipkow@55808: next nipkow@55808: case (Cons x xs ys) nipkow@55808: show ?case nipkow@55808: proof (cases "List.extract (op = x) ys") nipkow@55808: case None nipkow@55808: hence x: "x \ set ys" by (simp add: extract_None_iff) nipkow@55808: { nipkow@60515: assume "mset (x # xs) \# mset ys" nipkow@60495: from set_mset_mono[OF this] x have False by simp nipkow@55808: } note nle = this nipkow@55808: moreover nipkow@55808: { nipkow@60515: assume "mset (x # xs) <# mset ys" nipkow@60515: hence "mset (x # xs) \# mset ys" by auto nipkow@55808: from nle[OF this] have False . nipkow@55808: } nipkow@55808: ultimately show ?thesis using None by auto nipkow@55808: next nipkow@55808: case (Some res) nipkow@55808: obtain ys1 y ys2 where res: "res = (ys1,y,ys2)" by (cases res, auto) nipkow@55808: note Some = Some[unfolded res] nipkow@55808: from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp nipkow@60515: hence id: "mset ys = mset (ys1 @ ys2) + {#x#}" nipkow@55808: by (auto simp: ac_simps) nipkow@55808: show ?thesis unfolding ms_lesseq_impl.simps nipkow@55808: unfolding Some option.simps split nipkow@55808: unfolding id nipkow@55808: using Cons[of "ys1 @ ys2"] Mathias@60397: unfolding subset_mset_def subseteq_mset_def by auto nipkow@55808: qed nipkow@55808: qed nipkow@55808: nipkow@60515: lemma [code]: "mset xs \# mset ys \ ms_lesseq_impl xs ys \ None" nipkow@55808: using ms_lesseq_impl[of xs ys] by (cases "ms_lesseq_impl xs ys", auto) nipkow@55808: nipkow@60515: lemma [code]: "mset xs <# mset ys \ ms_lesseq_impl xs ys = Some True" nipkow@55808: using ms_lesseq_impl[of xs ys] by (cases "ms_lesseq_impl xs ys", auto) haftmann@51600: haftmann@51600: instantiation multiset :: (equal) equal haftmann@51600: begin haftmann@51600: haftmann@51600: definition nipkow@55808: [code del]: "HOL.equal A (B :: 'a multiset) \ A = B" nipkow@60515: lemma [code]: "HOL.equal (mset xs) (mset ys) \ ms_lesseq_impl xs ys = Some False" nipkow@55808: unfolding equal_multiset_def nipkow@55808: using ms_lesseq_impl[of xs ys] by (cases "ms_lesseq_impl xs ys", auto) haftmann@51600: haftmann@51600: instance wenzelm@60678: by standard (simp add: equal_multiset_def) wenzelm@60678: blanchet@37169: end haftmann@49388: wenzelm@60606: lemma [code]: "msetsum (mset xs) = listsum xs" haftmann@51600: by (induct xs) (simp_all add: add.commute) haftmann@51600: wenzelm@60606: lemma [code]: "msetprod (mset xs) = fold times xs 1" haftmann@51600: proof - nipkow@60515: have "\x. fold times xs x = msetprod (mset xs) * x" haftmann@51600: by (induct xs) (simp_all add: mult.assoc) haftmann@51600: then show ?thesis by simp haftmann@51600: qed haftmann@51600: wenzelm@60500: text \ haftmann@51600: Exercise for the casual reader: add implementations for @{const le_multiset} haftmann@51600: and @{const less_multiset} (multiset order). wenzelm@60500: \ wenzelm@60500: wenzelm@60500: text \Quickcheck generators\ haftmann@51600: haftmann@51600: definition (in term_syntax) wenzelm@61076: msetify :: "'a::typerep list \ (unit \ Code_Evaluation.term) haftmann@51600: \ 'a multiset \ (unit \ Code_Evaluation.term)" where nipkow@60515: [code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\} xs" haftmann@51600: haftmann@51600: notation fcomp (infixl "\>" 60) haftmann@51600: notation scomp (infixl "\\" 60) haftmann@51600: haftmann@51600: instantiation multiset :: (random) random haftmann@51600: begin haftmann@51600: haftmann@51600: definition haftmann@51600: "Quickcheck_Random.random i = Quickcheck_Random.random i \\ (\xs. Pair (msetify xs))" haftmann@51600: haftmann@51600: instance .. haftmann@51600: haftmann@51600: end haftmann@51600: haftmann@51600: no_notation fcomp (infixl "\>" 60) haftmann@51600: no_notation scomp (infixl "\\" 60) haftmann@51600: haftmann@51600: instantiation multiset :: (full_exhaustive) full_exhaustive haftmann@51600: begin haftmann@51600: haftmann@51600: definition full_exhaustive_multiset :: "('a multiset \ (unit \ term) \ (bool \ term list) option) \ natural \ (bool \ term list) option" haftmann@51600: where haftmann@51600: "full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (\xs. f (msetify xs)) i" haftmann@51600: haftmann@51600: instance .. haftmann@51600: haftmann@51600: end haftmann@51600: haftmann@51600: hide_const (open) msetify haftmann@51600: blanchet@55129: wenzelm@60500: subsection \BNF setup\ blanchet@55129: blanchet@57966: definition rel_mset where nipkow@60515: "rel_mset R X Y \ (\xs ys. mset xs = X \ mset ys = Y \ list_all2 R xs ys)" nipkow@60515: nipkow@60515: lemma mset_zip_take_Cons_drop_twice: blanchet@57966: assumes "length xs = length ys" "j \ length xs" nipkow@60515: shows "mset (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) = nipkow@60515: mset (zip xs ys) + {#(x, y)#}" wenzelm@60606: using assms blanchet@57966: proof (induct xs ys arbitrary: x y j rule: list_induct2) blanchet@57966: case Nil blanchet@57966: thus ?case blanchet@57966: by simp blanchet@55129: next blanchet@57966: case (Cons x xs y ys) blanchet@57966: thus ?case blanchet@57966: proof (cases "j = 0") blanchet@57966: case True blanchet@57966: thus ?thesis blanchet@57966: by simp blanchet@55129: next blanchet@57966: case False blanchet@57966: then obtain k where k: "j = Suc k" wenzelm@60678: by (cases j) simp blanchet@57966: hence "k \ length xs" blanchet@57966: using Cons.prems by auto nipkow@60515: hence "mset (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) = nipkow@60515: mset (zip xs ys) + {#(x, y)#}" blanchet@57966: by (rule Cons.hyps(2)) blanchet@57966: thus ?thesis blanchet@57966: unfolding k by (auto simp: add.commute union_lcomm) blanchet@58425: qed blanchet@55129: qed blanchet@55129: nipkow@60515: lemma ex_mset_zip_left: nipkow@60515: assumes "length xs = length ys" "mset xs' = mset xs" nipkow@60515: shows "\ys'. length ys' = length xs' \ mset (zip xs' ys') = mset (zip xs ys)" blanchet@58425: using assms blanchet@57966: proof (induct xs ys arbitrary: xs' rule: list_induct2) blanchet@57966: case Nil blanchet@57966: thus ?case blanchet@57966: by auto blanchet@57966: next blanchet@57966: case (Cons x xs y ys xs') blanchet@57966: obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" nipkow@60515: by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) blanchet@58425: blanchet@58425: def xsa \ "take j xs' @ drop (Suc j) xs'" nipkow@60515: have "mset xs' = {#x#} + mset xsa" blanchet@57966: unfolding xsa_def using j_len nth_j nipkow@58247: by (metis (no_types) ab_semigroup_add_class.add_ac(1) append_take_drop_id Cons_nth_drop_Suc nipkow@60515: mset.simps(2) union_code add.commute) nipkow@60515: hence ms_x: "mset xsa = mset xs" nipkow@60515: by (metis Cons.prems add.commute add_right_imp_eq mset.simps(2)) blanchet@57966: then obtain ysa where nipkow@60515: len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" blanchet@57966: using Cons.hyps(2) by blast blanchet@57966: blanchet@57966: def ys' \ "take j ysa @ y # drop j ysa" blanchet@57966: have xs': "xs' = take j xsa @ x # drop j xsa" blanchet@57966: using ms_x j_len nth_j Cons.prems xsa_def nipkow@58247: by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons nipkow@60515: length_drop size_mset) blanchet@57966: have j_len': "j \ length xsa" blanchet@57966: using j_len xs' xsa_def blanchet@57966: by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) blanchet@57966: have "length ys' = length xs'" blanchet@57966: unfolding ys'_def using Cons.prems len_a ms_x nipkow@60515: by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) nipkow@60515: moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" blanchet@57966: unfolding xs' ys'_def nipkow@60515: by (rule trans[OF mset_zip_take_Cons_drop_twice]) blanchet@57966: (auto simp: len_a ms_a j_len' add.commute) blanchet@57966: ultimately show ?case blanchet@57966: by blast blanchet@55129: qed blanchet@55129: blanchet@57966: lemma list_all2_reorder_left_invariance: nipkow@60515: assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs" nipkow@60515: shows "\ys'. list_all2 R xs' ys' \ mset ys' = mset ys" blanchet@57966: proof - blanchet@57966: have len: "length xs = length ys" blanchet@57966: using rel list_all2_conv_all_nth by auto blanchet@57966: obtain ys' where nipkow@60515: len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" nipkow@60515: using len ms_x by (metis ex_mset_zip_left) blanchet@57966: have "list_all2 R xs' ys'" nipkow@60515: using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) nipkow@60515: moreover have "mset ys' = mset ys" nipkow@60515: using len len' ms_xy map_snd_zip mset_map by metis blanchet@57966: ultimately show ?thesis blanchet@57966: by blast blanchet@55129: qed blanchet@55129: nipkow@60515: lemma ex_mset: "\xs. mset xs = X" nipkow@60515: by (induct X) (simp, metis mset.simps(2)) blanchet@55129: blanchet@55129: bnf "'a multiset" blanchet@57966: map: image_mset nipkow@60495: sets: set_mset blanchet@55129: bd: natLeq blanchet@55129: wits: "{#}" blanchet@57966: rel: rel_mset blanchet@57966: proof - blanchet@57966: show "image_mset id = id" blanchet@57966: by (rule image_mset.id) wenzelm@60606: show "image_mset (g \ f) = image_mset g \ image_mset f" for f g blanchet@59813: unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality) wenzelm@60606: show "(\z. z \ set_mset X \ f z = g z) \ image_mset f X = image_mset g X" for f g X wenzelm@60606: by (induct X) (simp_all (no_asm), nipkow@60495: metis One_nat_def Un_iff count_single mem_set_mset_iff set_mset_union zero_less_Suc) wenzelm@60606: show "set_mset \ image_mset f = op ` f \ set_mset" for f blanchet@57966: by auto blanchet@57966: show "card_order natLeq" blanchet@57966: by (rule natLeq_card_order) blanchet@57966: show "BNF_Cardinal_Arithmetic.cinfinite natLeq" blanchet@57966: by (rule natLeq_cinfinite) wenzelm@60606: show "ordLeq3 (card_of (set_mset X)) natLeq" for X blanchet@57966: by transfer blanchet@57966: (auto intro!: ordLess_imp_ordLeq simp: finite_iff_ordLess_natLeq[symmetric] multiset_def) wenzelm@60606: show "rel_mset R OO rel_mset S \ rel_mset (R OO S)" for R S blanchet@57966: unfolding rel_mset_def[abs_def] OO_def blanchet@57966: apply clarify wenzelm@60678: subgoal for X Z Y xs ys' ys zs wenzelm@60678: apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys]) wenzelm@60678: apply (auto intro: list_all2_trans) wenzelm@60678: done wenzelm@60606: done wenzelm@60606: show "rel_mset R = nipkow@60495: (BNF_Def.Grp {x. set_mset x \ {(x, y). R x y}} (image_mset fst))\\ OO wenzelm@60606: BNF_Def.Grp {x. set_mset x \ {(x, y). R x y}} (image_mset snd)" for R blanchet@57966: unfolding rel_mset_def[abs_def] BNF_Def.Grp_def OO_def blanchet@57966: apply (rule ext)+ blanchet@57966: apply auto nipkow@60515: apply (rule_tac x = "mset (zip xs ys)" in exI; auto) nipkow@60515: apply (metis list_all2_lengthD map_fst_zip mset_map) blanchet@57966: apply (auto simp: list_all2_iff)[1] nipkow@60515: apply (metis list_all2_lengthD map_snd_zip mset_map) blanchet@57966: apply (auto simp: list_all2_iff)[1] blanchet@57966: apply (rename_tac XY) nipkow@60515: apply (cut_tac X = XY in ex_mset) blanchet@57966: apply (erule exE) blanchet@57966: apply (rename_tac xys) blanchet@57966: apply (rule_tac x = "map fst xys" in exI) nipkow@60515: apply (auto simp: mset_map) blanchet@57966: apply (rule_tac x = "map snd xys" in exI) nipkow@60515: apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd) wenzelm@59997: done wenzelm@60606: show "z \ set_mset {#} \ False" for z blanchet@57966: by auto blanchet@57966: qed blanchet@57966: wenzelm@60606: inductive rel_mset' wenzelm@60606: where blanchet@57966: Zero[intro]: "rel_mset' R {#} {#}" blanchet@57966: | Plus[intro]: "\R a b; rel_mset' R M N\ \ rel_mset' R (M + {#a#}) (N + {#b#})" blanchet@57966: blanchet@57966: lemma rel_mset_Zero: "rel_mset R {#} {#}" blanchet@57966: unfolding rel_mset_def Grp_def by auto blanchet@55129: blanchet@55129: declare multiset.count[simp] blanchet@55129: declare Abs_multiset_inverse[simp] blanchet@55129: declare multiset.count_inverse[simp] blanchet@55129: declare union_preserves_multiset[simp] blanchet@55129: blanchet@57966: lemma rel_mset_Plus: wenzelm@60606: assumes ab: "R a b" wenzelm@60606: and MN: "rel_mset R M N" wenzelm@60606: shows "rel_mset R (M + {#a#}) (N + {#b#})" wenzelm@60606: proof - wenzelm@60606: have "\ya. image_mset fst y + {#a#} = image_mset fst ya \ wenzelm@60606: image_mset snd y + {#b#} = image_mset snd ya \ wenzelm@60606: set_mset ya \ {(x, y). R x y}" wenzelm@60606: if "R a b" and "set_mset y \ {(x, y). R x y}" for y wenzelm@60606: using that by (intro exI[of _ "y + {#(a,b)#}"]) auto blanchet@55129: thus ?thesis blanchet@55129: using assms blanchet@57966: unfolding multiset.rel_compp_Grp Grp_def by blast blanchet@55129: qed blanchet@55129: wenzelm@60606: lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \ rel_mset R M N" wenzelm@60678: by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus) blanchet@57966: wenzelm@60606: lemma rel_mset_size: "rel_mset R M N \ size M = size N" wenzelm@60678: unfolding multiset.rel_compp_Grp Grp_def by auto blanchet@55129: blanchet@55129: lemma multiset_induct2[case_names empty addL addR]: wenzelm@60678: assumes empty: "P {#} {#}" wenzelm@60678: and addL: "\M N a. P M N \ P (M + {#a#}) N" wenzelm@60678: and addR: "\M N a. P M N \ P M (N + {#a#})" wenzelm@60678: shows "P M N" blanchet@55129: apply(induct N rule: multiset_induct) blanchet@55129: apply(induct M rule: multiset_induct, rule empty, erule addL) blanchet@55129: apply(induct M rule: multiset_induct, erule addR, erule addR) blanchet@55129: done blanchet@55129: nipkow@59949: lemma multiset_induct2_size[consumes 1, case_names empty add]: wenzelm@60606: assumes c: "size M = size N" wenzelm@60606: and empty: "P {#} {#}" wenzelm@60606: and add: "\M N a b. P M N \ P (M + {#a#}) (N + {#b#})" wenzelm@60606: shows "P M N" wenzelm@60678: using c wenzelm@60678: proof (induct M arbitrary: N rule: measure_induct_rule[of size]) wenzelm@60606: case (less M) wenzelm@60606: show ?case blanchet@55129: proof(cases "M = {#}") blanchet@55129: case True hence "N = {#}" using less.prems by auto blanchet@55129: thus ?thesis using True empty by auto blanchet@55129: next blanchet@55129: case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split) blanchet@55129: have "N \ {#}" using False less.prems by auto blanchet@55129: then obtain N1 b where N: "N = N1 + {#b#}" by (metis multi_nonempty_split) nipkow@59949: have "size M1 = size N1" using less.prems unfolding M N by auto blanchet@55129: thus ?thesis using M N less.hyps add by auto blanchet@55129: qed blanchet@55129: qed blanchet@55129: blanchet@55129: lemma msed_map_invL: wenzelm@60606: assumes "image_mset f (M + {#a#}) = N" wenzelm@60606: shows "\N1. N = N1 + {#f a#} \ image_mset f M = N1" wenzelm@60606: proof - blanchet@55129: have "f a \# N" wenzelm@60606: using assms multiset.set_map[of f "M + {#a#}"] by auto blanchet@55129: then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis blanchet@57966: have "image_mset f M = N1" using assms unfolding N by simp blanchet@55129: thus ?thesis using N by blast blanchet@55129: qed blanchet@55129: blanchet@55129: lemma msed_map_invR: wenzelm@60606: assumes "image_mset f M = N + {#b#}" wenzelm@60606: shows "\M1 a. M = M1 + {#a#} \ f a = b \ image_mset f M1 = N" wenzelm@60606: proof - blanchet@55129: obtain a where a: "a \# M" and fa: "f a = b" wenzelm@60606: using multiset.set_map[of f M] unfolding assms wenzelm@60606: by (metis image_iff mem_set_mset_iff union_single_eq_member) blanchet@55129: then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis blanchet@57966: have "image_mset f M1 = N" using assms unfolding M fa[symmetric] by simp blanchet@55129: thus ?thesis using M fa by blast blanchet@55129: qed blanchet@55129: blanchet@55129: lemma msed_rel_invL: wenzelm@60606: assumes "rel_mset R (M + {#a#}) N" wenzelm@60606: shows "\N1 b. N = N1 + {#b#} \ R a b \ rel_mset R M N1" wenzelm@60606: proof - blanchet@57966: obtain K where KM: "image_mset fst K = M + {#a#}" wenzelm@60606: and KN: "image_mset snd K = N" and sK: "set_mset K \ {(a, b). R a b}" wenzelm@60606: using assms wenzelm@60606: unfolding multiset.rel_compp_Grp Grp_def by auto blanchet@55129: obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a" wenzelm@60606: and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto blanchet@57966: obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "image_mset snd K1 = N1" wenzelm@60606: using msed_map_invL[OF KN[unfolded K]] by auto blanchet@55129: have Rab: "R a (snd ab)" using sK a unfolding K by auto blanchet@57966: have "rel_mset R M N1" using sK K1M K1N1 wenzelm@60606: unfolding K multiset.rel_compp_Grp Grp_def by auto blanchet@55129: thus ?thesis using N Rab by auto blanchet@55129: qed blanchet@55129: blanchet@55129: lemma msed_rel_invR: wenzelm@60606: assumes "rel_mset R M (N + {#b#})" wenzelm@60606: shows "\M1 a. M = M1 + {#a#} \ R a b \ rel_mset R M1 N" wenzelm@60606: proof - blanchet@57966: obtain K where KN: "image_mset snd K = N + {#b#}" wenzelm@60606: and KM: "image_mset fst K = M" and sK: "set_mset K \ {(a, b). R a b}" wenzelm@60606: using assms wenzelm@60606: unfolding multiset.rel_compp_Grp Grp_def by auto blanchet@55129: obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b" wenzelm@60606: and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto blanchet@57966: obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "image_mset fst K1 = M1" wenzelm@60606: using msed_map_invL[OF KM[unfolded K]] by auto blanchet@55129: have Rab: "R (fst ab) b" using sK b unfolding K by auto blanchet@57966: have "rel_mset R M1 N" using sK K1N K1M1 wenzelm@60606: unfolding K multiset.rel_compp_Grp Grp_def by auto blanchet@55129: thus ?thesis using M Rab by auto blanchet@55129: qed blanchet@55129: blanchet@57966: lemma rel_mset_imp_rel_mset': wenzelm@60606: assumes "rel_mset R M N" wenzelm@60606: shows "rel_mset' R M N" nipkow@59949: using assms proof(induct M arbitrary: N rule: measure_induct_rule[of size]) blanchet@55129: case (less M) nipkow@59949: have c: "size M = size N" using rel_mset_size[OF less.prems] . blanchet@55129: show ?case blanchet@55129: proof(cases "M = {#}") blanchet@55129: case True hence "N = {#}" using c by simp blanchet@57966: thus ?thesis using True rel_mset'.Zero by auto blanchet@55129: next blanchet@55129: case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split) blanchet@57966: obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "rel_mset R M1 N1" wenzelm@60606: using msed_rel_invL[OF less.prems[unfolded M]] by auto blanchet@57966: have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp blanchet@57966: thus ?thesis using rel_mset'.Plus[of R a b, OF R] unfolding M N by simp blanchet@55129: qed blanchet@55129: qed blanchet@55129: wenzelm@60606: lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N" wenzelm@60678: using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto blanchet@57966: hoelzl@60613: text \The main end product for @{const rel_mset}: inductive characterization:\ wenzelm@61337: lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] = wenzelm@60606: rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]] blanchet@55129: blanchet@56656: wenzelm@60500: subsection \Size setup\ blanchet@56656: blanchet@57966: lemma multiset_size_o_map: "size_multiset g \ image_mset f = size_multiset (g \ f)" wenzelm@60678: apply (rule ext) wenzelm@60678: subgoal for x by (induct x) auto wenzelm@60678: done blanchet@56656: wenzelm@60500: setup \ wenzelm@60606: BNF_LFP_Size.register_size_global @{type_name multiset} @{const_name size_multiset} wenzelm@60606: @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single wenzelm@60606: size_union} wenzelm@60606: @{thms multiset_size_o_map} wenzelm@60500: \ blanchet@56656: blanchet@56656: hide_const (open) wcount blanchet@56656: blanchet@55129: end