wenzelm@10249: (* Title: HOL/Library/Multiset.thy paulson@15072: Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker wenzelm@10249: *) wenzelm@10249: haftmann@34943: header {* (Finite) multisets *} wenzelm@10249: nipkow@15131: theory Multiset bulwahn@46237: imports Main DAList nipkow@15131: begin wenzelm@10249: wenzelm@10249: subsection {* The type of multisets *} wenzelm@10249: wenzelm@45694: definition "multiset = {f :: 'a => nat. finite {x. f x > 0}}" wenzelm@45694: wenzelm@45694: typedef (open) '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: haftmann@28708: abbreviation Melem :: "'a => 'a multiset => bool" ("(_/ :# _)" [50, 51] 50) where kleing@25610: "a :# M == 0 < count M a" kleing@25610: wenzelm@26145: notation (xsymbols) wenzelm@26145: Melem (infix "\#" 50) wenzelm@10249: nipkow@39302: lemma multiset_eq_iff: haftmann@34943: "M = N \ (\a. count M a = count N a)" nipkow@39302: by (simp only: count_inject [symmetric] fun_eq_iff) haftmann@34943: nipkow@39302: lemma multiset_eqI: haftmann@34943: "(\x. count A x = count B x) \ A = B" nipkow@39302: using multiset_eq_iff by auto haftmann@34943: haftmann@34943: text {* haftmann@34943: \medskip Preservation of the representing set @{term multiset}. haftmann@34943: *} haftmann@34943: haftmann@34943: lemma const0_in_multiset: haftmann@34943: "(\a. 0) \ multiset" haftmann@34943: by (simp add: multiset_def) haftmann@34943: haftmann@34943: lemma only1_in_multiset: haftmann@34943: "(\b. if b = a then n else 0) \ multiset" haftmann@34943: by (simp add: multiset_def) haftmann@34943: haftmann@34943: lemma union_preserves_multiset: haftmann@34943: "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: haftmann@34943: subsection {* Representing multisets *} haftmann@34943: haftmann@34943: 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: bulwahn@47429: 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: huffman@48008: instance huffman@48008: by default (transfer, simp add: fun_eq_iff)+ haftmann@25571: haftmann@25571: end wenzelm@10249: bulwahn@47429: 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@26176: "_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: haftmann@34943: subsection {* Basic operations *} wenzelm@10249: wenzelm@10249: 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@10249: subsubsection {* Difference *} wenzelm@10249: haftmann@49388: instantiation multiset :: (type) comm_monoid_diff haftmann@34943: begin haftmann@34943: bulwahn@47429: lift_definition minus_multiset :: "'a multiset => 'a multiset => 'a multiset" is "\ M N. \a. M a - N a" bulwahn@47429: by (rule diff_preserves_multiset) bulwahn@47429: haftmann@49388: instance haftmann@49388: by default (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 = {#}" nipkow@39302: by(simp add: multiset_eq_iff) nipkow@36903: nipkow@36903: lemma diff_cancel[simp]: "A - A = {#}" nipkow@39302: by (rule multiset_eqI) simp wenzelm@10249: nipkow@36903: lemma diff_union_cancelR [simp]: "M + N - N = (M::'a multiset)" nipkow@39302: by(simp add: multiset_eq_iff) wenzelm@10249: nipkow@36903: lemma diff_union_cancelL [simp]: "N + M - N = (M::'a multiset)" nipkow@39302: by(simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma insert_DiffM: haftmann@34943: "x \# M \ {#x#} + (M - {#x#}) = M" nipkow@39302: by (clarsimp simp: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma insert_DiffM2 [simp]: haftmann@34943: "x \# M \ M - {#x#} + {#x#} = M" nipkow@39302: by (clarsimp simp: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma diff_right_commute: haftmann@34943: "(M::'a multiset) - N - Q = M - Q - N" nipkow@39302: by (auto simp add: multiset_eq_iff) nipkow@36903: nipkow@36903: lemma diff_add: nipkow@36903: "(M::'a multiset) - (N + Q) = M - N - Q" nipkow@39302: by (simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma diff_union_swap: haftmann@34943: "a \ b \ M - {#a#} + {#b#} = M + {#b#} - {#a#}" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma diff_union_single_conv: haftmann@34943: "a \# J \ I + J - {#a#} = I + (J - {#a#})" nipkow@39302: by (simp add: multiset_eq_iff) bulwahn@26143: wenzelm@10249: haftmann@34943: 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: haftmann@34943: lemma diff_single_trivial: haftmann@34943: "\ x \# M \ M - {#x#} = M" nipkow@39302: by (auto simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma diff_single_eq_union: haftmann@34943: "x \# M \ M - {#x#} = N \ M = N + {#x#}" haftmann@34943: by auto haftmann@34943: haftmann@34943: lemma union_single_eq_diff: haftmann@34943: "M + {#x#} = N \ M = N - {#x#}" haftmann@34943: by (auto dest: sym) haftmann@34943: haftmann@34943: lemma union_single_eq_member: haftmann@34943: "M + {#x#} = N \ x \# N" haftmann@34943: by auto haftmann@34943: haftmann@34943: lemma union_is_single: wenzelm@46730: "M + N = {#a#} \ M = {#a#} \ N={#} \ M = {#} \ N = {#a#}" (is "?lhs = ?rhs") wenzelm@46730: proof haftmann@34943: assume ?rhs then show ?lhs by auto haftmann@34943: next wenzelm@46730: assume ?lhs then show ?rhs wenzelm@46730: by (simp add: multiset_eq_iff split:if_splits) (metis add_is_1) haftmann@34943: qed haftmann@34943: haftmann@34943: lemma single_is_union: haftmann@34943: "{#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: haftmann@34943: "M + {#a#} = N + {#b#} \ M = N \ a = b \ M = N - {#a#} + {#b#} \ N = M - {#b#} + {#a#}" (is "?lhs = ?rhs") nipkow@44890: (* shorter: by (simp add: multiset_eq_iff) fastforce *) haftmann@34943: proof haftmann@34943: assume ?rhs then show ?lhs haftmann@34943: by (auto simp add: add_assoc add_commute [of "{#b#}"]) haftmann@34943: (drule sym, simp add: add_assoc [symmetric]) haftmann@34943: next haftmann@34943: assume ?lhs haftmann@34943: show ?rhs haftmann@34943: proof (cases "a = b") haftmann@34943: case True with `?lhs` show ?thesis by simp haftmann@34943: next haftmann@34943: case False haftmann@34943: from `?lhs` have "a \# N + {#b#}" by (rule union_single_eq_member) haftmann@34943: with False have "a \# N" by auto haftmann@34943: 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: haftmann@34943: 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: haftmann@34943: haftmann@34943: subsubsection {* Pointwise ordering induced by count *} haftmann@34943: haftmann@35268: instantiation multiset :: (type) ordered_ab_semigroup_add_imp_le haftmann@35268: begin haftmann@35268: bulwahn@47429: lift_definition less_eq_multiset :: "'a multiset \ 'a multiset \ bool" is "\ A B. (\a. A a \ B a)" bulwahn@47429: by simp bulwahn@47429: lemmas mset_le_def = less_eq_multiset_def haftmann@34943: haftmann@35268: definition less_multiset :: "'a multiset \ 'a multiset \ bool" where haftmann@35268: mset_less_def: "(A::'a multiset) < B \ A \ B \ A \ B" haftmann@34943: wenzelm@46921: instance wenzelm@46921: by default (auto simp add: mset_le_def mset_less_def multiset_eq_iff intro: order_trans antisym) haftmann@35268: haftmann@35268: end haftmann@34943: haftmann@34943: lemma mset_less_eqI: haftmann@35268: "(\x. count A x \ count B x) \ A \ B" haftmann@34943: by (simp add: mset_le_def) haftmann@34943: haftmann@35268: lemma mset_le_exists_conv: haftmann@35268: "(A::'a multiset) \ B \ (\C. B = A + C)" haftmann@34943: apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI) nipkow@39302: apply (auto intro: multiset_eq_iff [THEN iffD2]) haftmann@34943: done haftmann@34943: haftmann@35268: lemma mset_le_mono_add_right_cancel [simp]: haftmann@35268: "(A::'a multiset) + C \ B + C \ A \ B" haftmann@35268: by (fact add_le_cancel_right) haftmann@34943: haftmann@35268: lemma mset_le_mono_add_left_cancel [simp]: haftmann@35268: "C + (A::'a multiset) \ C + B \ A \ B" haftmann@35268: by (fact add_le_cancel_left) haftmann@35268: haftmann@35268: lemma mset_le_mono_add: haftmann@35268: "(A::'a multiset) \ B \ C \ D \ A + C \ B + D" haftmann@35268: by (fact add_mono) haftmann@34943: haftmann@35268: lemma mset_le_add_left [simp]: haftmann@35268: "(A::'a multiset) \ A + B" haftmann@35268: unfolding mset_le_def by auto haftmann@35268: haftmann@35268: lemma mset_le_add_right [simp]: haftmann@35268: "B \ (A::'a multiset) + B" haftmann@35268: unfolding mset_le_def by auto haftmann@34943: haftmann@35268: lemma mset_le_single: haftmann@35268: "a :# B \ {#a#} \ B" haftmann@35268: by (simp add: mset_le_def) haftmann@34943: haftmann@35268: lemma multiset_diff_union_assoc: haftmann@35268: "C \ B \ (A::'a multiset) + B - C = A + (B - C)" nipkow@39302: by (simp add: multiset_eq_iff mset_le_def) haftmann@34943: haftmann@34943: lemma mset_le_multiset_union_diff_commute: nipkow@36867: "B \ A \ (A::'a multiset) - B + C = A + C - B" nipkow@39302: by (simp add: multiset_eq_iff mset_le_def) haftmann@34943: nipkow@39301: lemma diff_le_self[simp]: "(M::'a multiset) - N \ M" nipkow@39301: by(simp add: mset_le_def) nipkow@39301: haftmann@35268: lemma mset_lessD: "A < B \ x \# A \ x \# B" haftmann@34943: apply (clarsimp simp: mset_le_def mset_less_def) haftmann@34943: apply (erule_tac x=x in allE) haftmann@34943: apply auto haftmann@34943: done haftmann@34943: haftmann@35268: lemma mset_leD: "A \ B \ x \# A \ x \# B" haftmann@34943: apply (clarsimp simp: mset_le_def mset_less_def) haftmann@34943: apply (erule_tac x = x in allE) haftmann@34943: apply auto haftmann@34943: done haftmann@34943: haftmann@35268: lemma mset_less_insertD: "(A + {#x#} < B) \ (x \# B \ A < B)" haftmann@34943: apply (rule conjI) haftmann@34943: apply (simp add: mset_lessD) haftmann@34943: apply (clarsimp simp: mset_le_def mset_less_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: haftmann@35268: lemma mset_le_insertD: "(A + {#x#} \ B) \ (x \# B \ A \ B)" haftmann@34943: apply (rule conjI) haftmann@34943: apply (simp add: mset_leD) haftmann@34943: apply (force simp: mset_le_def mset_less_def split: split_if_asm) haftmann@34943: done haftmann@34943: haftmann@35268: lemma mset_less_of_empty[simp]: "A < {#} \ False" nipkow@39302: by (auto simp add: mset_less_def mset_le_def multiset_eq_iff) haftmann@34943: haftmann@35268: lemma multi_psub_of_add_self[simp]: "A < A + {#x#}" haftmann@35268: by (auto simp: mset_le_def mset_less_def) haftmann@34943: haftmann@35268: lemma multi_psub_self[simp]: "(A::'a multiset) < A = False" haftmann@35268: by simp haftmann@34943: haftmann@34943: lemma mset_less_add_bothsides: haftmann@35268: "T + {#x#} < S + {#x#} \ T < S" haftmann@35268: by (fact add_less_imp_less_right) haftmann@35268: haftmann@35268: lemma mset_less_empty_nonempty: haftmann@35268: "{#} < S \ S \ {#}" haftmann@35268: by (auto simp: mset_le_def mset_less_def) haftmann@35268: haftmann@35268: lemma mset_less_diff_self: haftmann@35268: "c \# B \ B - {#c#} < B" nipkow@39302: by (auto simp: mset_le_def mset_less_def multiset_eq_iff) haftmann@35268: haftmann@35268: haftmann@35268: subsubsection {* Intersection *} haftmann@35268: haftmann@35268: instantiation multiset :: (type) semilattice_inf haftmann@35268: begin haftmann@35268: haftmann@35268: definition inf_multiset :: "'a multiset \ 'a multiset \ 'a multiset" where haftmann@35268: multiset_inter_def: "inf_multiset A B = A - (A - B)" haftmann@35268: wenzelm@46921: instance wenzelm@46921: proof - haftmann@35268: have aux: "\m n q :: nat. m \ n \ m \ q \ m \ n - (n - q)" by arith wenzelm@46921: show "OFCLASS('a multiset, semilattice_inf_class)" wenzelm@46921: by default (auto simp add: multiset_inter_def mset_le_def aux) haftmann@35268: qed haftmann@35268: haftmann@35268: end haftmann@35268: haftmann@35268: abbreviation multiset_inter :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "#\" 70) where haftmann@35268: "multiset_inter \ inf" haftmann@34943: haftmann@41069: lemma multiset_inter_count [simp]: haftmann@35268: "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: haftmann@35268: haftmann@41069: subsubsection {* Filter (with comprehension syntax) *} haftmann@41069: haftmann@41069: text {* Multiset comprehension *} haftmann@41069: bulwahn@47429: lift_definition filter :: "('a \ bool) \ 'a multiset \ 'a multiset" is "\P M. \x. if P x then M x else 0" bulwahn@47429: by (rule filter_preserves_multiset) haftmann@35268: haftmann@41069: hide_const (open) filter haftmann@35268: haftmann@41069: lemma count_filter [simp]: haftmann@41069: "count (Multiset.filter P M) a = (if P a then count M a else 0)" bulwahn@47429: by (simp add: filter.rep_eq) haftmann@41069: haftmann@41069: lemma filter_empty [simp]: haftmann@41069: "Multiset.filter P {#} = {#}" nipkow@39302: by (rule multiset_eqI) simp haftmann@35268: haftmann@41069: lemma filter_single [simp]: haftmann@41069: "Multiset.filter P {#x#} = (if P x then {#x#} else {#})" haftmann@41069: by (rule multiset_eqI) simp haftmann@41069: haftmann@41069: lemma filter_union [simp]: haftmann@41069: "Multiset.filter P (M + N) = Multiset.filter P M + Multiset.filter P N" nipkow@39302: by (rule multiset_eqI) simp haftmann@35268: haftmann@41069: lemma filter_diff [simp]: haftmann@41069: "Multiset.filter P (M - N) = Multiset.filter P M - Multiset.filter P N" haftmann@41069: by (rule multiset_eqI) simp haftmann@41069: haftmann@41069: lemma filter_inter [simp]: haftmann@41069: "Multiset.filter P (M #\ N) = Multiset.filter P M #\ Multiset.filter P N" nipkow@39302: by (rule multiset_eqI) simp wenzelm@10249: 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 haftmann@41069: "{#x \# M. P#}" == "CONST Multiset.filter (\x. P) M" haftmann@41069: wenzelm@10249: wenzelm@10249: subsubsection {* Set of elements *} wenzelm@10249: haftmann@34943: definition set_of :: "'a multiset => 'a set" where haftmann@34943: "set_of M = {x. x :# M}" haftmann@34943: wenzelm@17161: lemma set_of_empty [simp]: "set_of {#} = {}" nipkow@26178: by (simp add: set_of_def) wenzelm@10249: wenzelm@17161: lemma set_of_single [simp]: "set_of {#b#} = {b}" nipkow@26178: by (simp add: set_of_def) wenzelm@10249: wenzelm@17161: lemma set_of_union [simp]: "set_of (M + N) = set_of M \ set_of N" nipkow@26178: by (auto simp add: set_of_def) wenzelm@10249: wenzelm@17161: lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})" nipkow@39302: by (auto simp add: set_of_def multiset_eq_iff) wenzelm@10249: wenzelm@17161: lemma mem_set_of_iff [simp]: "(x \ set_of M) = (x :# M)" nipkow@26178: by (auto simp add: set_of_def) nipkow@26016: haftmann@41069: lemma set_of_filter [simp]: "set_of {# x:#M. P x #} = set_of M \ {x. P x}" nipkow@26178: by (auto simp add: set_of_def) wenzelm@10249: haftmann@34943: lemma finite_set_of [iff]: "finite (set_of M)" haftmann@34943: using count [of M] by (simp add: multiset_def set_of_def) haftmann@34943: bulwahn@46756: lemma finite_Collect_mem [iff]: "finite {x. x :# M}" bulwahn@46756: unfolding set_of_def[symmetric] by simp wenzelm@10249: wenzelm@10249: subsubsection {* Size *} wenzelm@10249: haftmann@34943: instantiation multiset :: (type) size haftmann@34943: begin haftmann@34943: haftmann@34943: definition size_def: haftmann@34943: "size M = setsum (count M) (set_of M)" haftmann@34943: haftmann@34943: instance .. haftmann@34943: haftmann@34943: end haftmann@34943: haftmann@28708: lemma size_empty [simp]: "size {#} = 0" nipkow@26178: by (simp add: size_def) wenzelm@10249: haftmann@28708: lemma size_single [simp]: "size {#b#} = 1" nipkow@26178: by (simp add: size_def) wenzelm@10249: wenzelm@17161: lemma setsum_count_Int: nipkow@26178: "finite A ==> setsum (count N) (A \ set_of N) = setsum (count N) A" nipkow@26178: apply (induct rule: finite_induct) nipkow@26178: apply simp nipkow@26178: apply (simp add: Int_insert_left set_of_def) nipkow@26178: done wenzelm@10249: haftmann@28708: lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" nipkow@26178: apply (unfold size_def) nipkow@26178: apply (subgoal_tac "count (M + N) = (\a. count M a + count N a)") nipkow@26178: prefer 2 nipkow@26178: apply (rule ext, simp) nipkow@26178: apply (simp (no_asm_simp) add: setsum_Un_nat setsum_addf setsum_count_Int) nipkow@26178: apply (subst Int_commute) nipkow@26178: apply (simp (no_asm_simp) add: setsum_count_Int) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" nipkow@39302: by (auto simp add: size_def multiset_eq_iff) nipkow@26016: nipkow@26016: lemma nonempty_has_size: "(S \ {#}) = (0 < size S)" nipkow@26178: by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) wenzelm@10249: wenzelm@17161: lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \a. a :# M" nipkow@26178: apply (unfold size_def) nipkow@26178: apply (drule setsum_SucD) nipkow@26178: apply auto nipkow@26178: done wenzelm@10249: 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: nipkow@26016: nipkow@26016: 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#}" huffman@48009: 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: kleing@25610: lemma multiset_cases [cases type, case_names empty add]: nipkow@26178: assumes em: "M = {#} \ P" nipkow@26178: assumes add: "\N x. M = N + {#x#} \ P" nipkow@26178: shows "P" huffman@48009: using assms by (induct M) simp_all kleing@25610: kleing@25610: lemma multi_member_split: "x \# M \ \A. M = A + {#x#}" huffman@48009: by (rule_tac x="M - {#x#}" in exI, simp) 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: nipkow@26033: 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: haftmann@35268: 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) haftmann@34943: 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) haftmann@35268: have IH: "\B. S < B \ size S < size B" by fact haftmann@35268: have SxsubT: "S + {#x#} < T" by fact haftmann@35268: then have "x \# T" and "S < T" by (auto dest: mset_less_insertD) haftmann@34943: then obtain T' where T: "T = T' + {#x#}" haftmann@34943: by (blast dest: multi_member_split) haftmann@35268: 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: haftmann@34943: subsubsection {* Strong induction and subset induction for multisets *} haftmann@34943: haftmann@34943: text {* Well-foundedness of proper subset operator: *} haftmann@34943: haftmann@34943: text {* proper multiset subset *} haftmann@34943: haftmann@34943: definition haftmann@34943: mset_less_rel :: "('a multiset * 'a multiset) set" where haftmann@35268: "mset_less_rel = {(A,B). A < B}" wenzelm@10249: haftmann@34943: lemma multiset_add_sub_el_shuffle: haftmann@34943: assumes "c \# B" and "b \ c" haftmann@34943: shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}" haftmann@34943: proof - haftmann@34943: from `c \# B` obtain A where B: "B = A + {#c#}" haftmann@34943: by (blast dest: multi_member_split) haftmann@34943: have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp haftmann@34943: then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" haftmann@34943: by (simp add: add_ac) haftmann@34943: then show ?thesis using B by simp haftmann@34943: qed haftmann@34943: haftmann@34943: lemma wf_mset_less_rel: "wf mset_less_rel" haftmann@34943: apply (unfold mset_less_rel_def) 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: text {* The induction rules: *} haftmann@34943: haftmann@34943: lemma full_multiset_induct [case_names less]: haftmann@35268: assumes ih: "\B. \(A::'a multiset). A < B \ P A \ P B" haftmann@34943: shows "P B" haftmann@34943: apply (rule wf_mset_less_rel [THEN wf_induct]) haftmann@34943: apply (rule ih, auto simp: mset_less_rel_def) haftmann@34943: done haftmann@34943: haftmann@34943: lemma multi_subset_induct [consumes 2, case_names empty add]: haftmann@35268: assumes "F \ A" haftmann@34943: and empty: "P {#}" haftmann@34943: and insert: "\a F. a \# A \ P F \ P (F + {#a#})" haftmann@34943: shows "P F" haftmann@34943: proof - haftmann@35268: 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 haftmann@35268: 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) haftmann@35268: 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: huffman@48023: subsection {* The fold combinator *} huffman@48023: haftmann@49822: definition fold :: "('a \ 'b \ 'b) \ 'b \ 'a multiset \ 'b" huffman@48023: where haftmann@49822: "fold f s M = Finite_Set.fold (\x. f x ^^ count M x) s (set_of M)" huffman@48023: haftmann@49822: lemma fold_mset_empty [simp]: haftmann@49822: "fold f s {#} = s" haftmann@49822: by (simp add: fold_def) huffman@48023: huffman@48023: context comp_fun_commute huffman@48023: begin huffman@48023: haftmann@49822: lemma fold_mset_insert: haftmann@49822: "fold f s (M + {#x#}) = f x (fold 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 haftmann@49822: proof (cases "x \ set_of M") haftmann@49822: case False haftmann@49822: then have *: "count (M + {#x#}) x = 1" by simp haftmann@49822: from False have "Finite_Set.fold (\y. f y ^^ count (M + {#x#}) y) s (set_of M) = haftmann@49822: Finite_Set.fold (\y. f y ^^ count M y) s (set_of M)" haftmann@49822: by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) haftmann@49822: with False * show ?thesis haftmann@49822: by (simp add: fold_def del: count_union) huffman@48023: next haftmann@49822: case True haftmann@49822: def N \ "set_of M - {x}" haftmann@49822: from N_def True have *: "set_of 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) haftmann@49822: with * show ?thesis by (simp add: fold_def del: count_union) simp huffman@48023: qed huffman@48023: qed huffman@48023: haftmann@49822: corollary fold_mset_single [simp]: haftmann@49822: "fold f s {#x#} = f x s" haftmann@49822: proof - haftmann@49822: have "fold 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: haftmann@49822: lemma fold_mset_fun_comm: haftmann@49822: "f x (fold f s M) = fold f (f x s) M" haftmann@49822: by (induct M) (simp_all add: fold_mset_insert fun_left_comm) huffman@48023: huffman@48023: lemma fold_mset_union [simp]: haftmann@49822: "fold f s (M + N) = fold f (fold 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@49822: by (simp add: add_ac) haftmann@49822: with add show ?case by (simp add: fold_mset_insert fold_mset_fun_comm) huffman@48023: qed huffman@48023: huffman@48023: lemma fold_mset_fusion: huffman@48023: assumes "comp_fun_commute g" haftmann@49822: shows "(\x y. h (g x y) = f x (h y)) \ h (fold g w A) = fold f (h w) A" (is "PROP ?P") huffman@48023: proof - huffman@48023: interpret comp_fun_commute g by (fact assms) huffman@48023: show "PROP ?P" by (induct A) auto huffman@48023: qed huffman@48023: huffman@48023: end huffman@48023: huffman@48023: text {* huffman@48023: A note on code generation: When defining some function containing a haftmann@49822: subterm @{term "fold F"}, code generation is not automatic. When huffman@48023: interpreting locale @{text left_commutative} with @{text F}, the haftmann@49822: would be code thms for @{const fold} become thms like haftmann@49822: @{term "fold F z {#} = z"} where @{text F} is not a pattern but huffman@48023: contains defined symbols, i.e.\ is not a code thm. Hence a separate huffman@48023: constant with its own code thms needs to be introduced for @{text huffman@48023: F}. See the image operator below. huffman@48023: *} huffman@48023: huffman@48023: huffman@48023: subsection {* Image *} huffman@48023: huffman@48023: definition image_mset :: "('a \ 'b) \ 'a multiset \ 'b multiset" where haftmann@49822: "image_mset f = fold (plus o single o f) {#}" huffman@48023: huffman@48023: interpretation image_fun_commute: comp_fun_commute "op + o single o f" for f huffman@48023: proof qed (simp add: add_ac fun_eq_iff) huffman@48023: huffman@48023: lemma image_mset_empty [simp]: "image_mset f {#} = {#}" huffman@48023: by (simp add: image_mset_def) huffman@48023: huffman@48023: lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}" huffman@48023: by (simp add: image_mset_def) huffman@48023: huffman@48023: lemma image_mset_insert: huffman@48023: "image_mset f (M + {#a#}) = image_mset f M + {#f a#}" huffman@48023: by (simp add: image_mset_def add_ac) huffman@48023: huffman@48023: lemma image_mset_union [simp]: huffman@48023: "image_mset f (M+N) = image_mset f M + image_mset f N" huffman@48023: apply (induct N) huffman@48023: apply simp huffman@48023: apply (simp add: add_assoc [symmetric] image_mset_insert) huffman@48023: done huffman@48023: huffman@48040: lemma set_of_image_mset [simp]: "set_of (image_mset f M) = image f (set_of M)" huffman@48040: by (induct M) simp_all huffman@48040: huffman@48023: lemma size_image_mset [simp]: "size (image_mset f M) = size M" huffman@48023: by (induct M) simp_all huffman@48023: huffman@48023: lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \ M = {#}" huffman@48023: by (cases M) auto huffman@48023: huffman@48023: syntax huffman@48023: "_comprehension1_mset" :: "'a \ 'b \ 'b multiset \ 'a multiset" huffman@48023: ("({#_/. _ :# _#})") huffman@48023: translations huffman@48023: "{#e. x:#M#}" == "CONST image_mset (%x. e) M" huffman@48023: huffman@48023: syntax huffman@48023: "_comprehension2_mset" :: "'a \ 'b \ 'b multiset \ bool \ 'a multiset" huffman@48023: ("({#_/ | _ :# _./ _#})") huffman@48023: translations huffman@48023: "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}" huffman@48023: huffman@48023: text {* huffman@48023: This allows to write not just filters like @{term "{#x:#M. x 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: haftmann@49717: declare image_mset.identity [simp] haftmann@49717: huffman@48023: haftmann@34943: subsection {* Alternative representations *} haftmann@34943: haftmann@34943: subsubsection {* Lists *} haftmann@34943: haftmann@34943: primrec multiset_of :: "'a list \ 'a multiset" where haftmann@34943: "multiset_of [] = {#}" | haftmann@34943: "multiset_of (a # x) = multiset_of x + {# a #}" haftmann@34943: haftmann@37107: lemma in_multiset_in_set: haftmann@37107: "x \# multiset_of xs \ x \ set xs" haftmann@37107: by (induct xs) simp_all haftmann@37107: haftmann@37107: lemma count_multiset_of: haftmann@37107: "count (multiset_of xs) x = length (filter (\y. x = y) xs)" haftmann@37107: by (induct xs) simp_all haftmann@37107: haftmann@34943: lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])" haftmann@34943: by (induct x) auto haftmann@34943: haftmann@34943: lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])" haftmann@34943: by (induct x) auto haftmann@34943: haftmann@40950: lemma set_of_multiset_of[simp]: "set_of (multiset_of x) = set x" haftmann@34943: by (induct x) auto haftmann@34943: haftmann@34943: lemma mem_set_multiset_eq: "x \ set xs = (x :# multiset_of xs)" haftmann@34943: by (induct xs) auto haftmann@34943: huffman@48012: lemma size_multiset_of [simp]: "size (multiset_of xs) = length xs" huffman@48012: by (induct xs) simp_all huffman@48012: haftmann@34943: lemma multiset_of_append [simp]: haftmann@34943: "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" haftmann@34943: by (induct xs arbitrary: ys) (auto simp: add_ac) haftmann@34943: haftmann@40303: lemma multiset_of_filter: haftmann@40303: "multiset_of (filter P xs) = {#x :# multiset_of xs. P x #}" haftmann@40303: by (induct xs) simp_all haftmann@40303: haftmann@40950: lemma multiset_of_rev [simp]: haftmann@40950: "multiset_of (rev xs) = multiset_of xs" haftmann@40950: by (induct xs) simp_all haftmann@40950: haftmann@34943: lemma surj_multiset_of: "surj multiset_of" 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: haftmann@34943: lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}" haftmann@34943: by (induct x) auto haftmann@34943: haftmann@34943: lemma distinct_count_atmost_1: haftmann@34943: "distinct x = (! a. count (multiset_of x) a = (if a \ set x then 1 else 0))" haftmann@34943: apply (induct x, simp, rule iffI, simp_all) haftmann@34943: apply (rule conjI) haftmann@34943: apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) haftmann@34943: apply (erule_tac x = a in allE, simp, clarify) haftmann@34943: apply (erule_tac x = aa in allE, simp) haftmann@34943: done haftmann@34943: haftmann@34943: lemma multiset_of_eq_setD: haftmann@34943: "multiset_of xs = multiset_of ys \ set xs = set ys" nipkow@39302: by (rule) (auto simp add:multiset_eq_iff set_count_greater_0) haftmann@34943: haftmann@34943: lemma set_eq_iff_multiset_of_eq_distinct: haftmann@34943: "distinct x \ distinct y \ haftmann@34943: (set x = set y) = (multiset_of x = multiset_of y)" nipkow@39302: by (auto simp: multiset_eq_iff distinct_count_atmost_1) haftmann@34943: haftmann@34943: lemma set_eq_iff_multiset_of_remdups_eq: haftmann@34943: "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" haftmann@34943: apply (rule iffI) haftmann@34943: apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) haftmann@34943: apply (drule distinct_remdups [THEN distinct_remdups haftmann@34943: [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]]) haftmann@34943: apply simp haftmann@34943: done haftmann@34943: haftmann@34943: lemma multiset_of_compl_union [simp]: haftmann@34943: "multiset_of [x\xs. P x] + multiset_of [x\xs. \P x] = multiset_of xs" haftmann@34943: by (induct xs) (auto simp: add_ac) haftmann@34943: haftmann@41069: lemma count_multiset_of_length_filter: haftmann@39533: "count (multiset_of xs) x = length (filter (\y. x = y) xs)" haftmann@39533: by (induct xs) auto haftmann@34943: haftmann@34943: lemma nth_mem_multiset_of: "i < length ls \ (ls ! i) :# multiset_of ls" haftmann@34943: apply (induct ls arbitrary: i) haftmann@34943: apply simp haftmann@34943: apply (case_tac i) haftmann@34943: apply auto haftmann@34943: done haftmann@34943: nipkow@36903: lemma multiset_of_remove1[simp]: nipkow@36903: "multiset_of (remove1 a xs) = multiset_of xs - {#a#}" nipkow@39302: by (induct xs) (auto simp add: multiset_eq_iff) haftmann@34943: haftmann@34943: lemma multiset_of_eq_length: haftmann@37107: assumes "multiset_of xs = multiset_of ys" haftmann@37107: shows "length xs = length ys" huffman@48012: using assms by (metis size_multiset_of) haftmann@34943: haftmann@39533: lemma multiset_of_eq_length_filter: haftmann@39533: assumes "multiset_of xs = multiset_of ys" haftmann@39533: shows "length (filter (\x. z = x) xs) = length (filter (\y. z = y) ys)" huffman@48012: using assms by (metis count_multiset_of) 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" haftmann@45989: and equiv: "multiset_of xs = multiset_of ys" haftmann@49822: shows "List.fold f xs = List.fold f ys" wenzelm@46921: using f equiv [symmetric] wenzelm@46921: proof (induct xs arbitrary: ys) haftmann@45989: case Nil then show ?case by simp haftmann@45989: next haftmann@45989: case (Cons x xs) haftmann@45989: then have *: "set ys = set (x # xs)" by (blast dest: multiset_of_eq_setD) haftmann@45989: 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: *) haftmann@45989: moreover from * have "x \ set ys" by simp haftmann@49822: ultimately have "List.fold f ys = List.fold f (remove1 x ys) \ f x" by (fact fold_remove1_split) haftmann@49822: moreover from Cons.prems have "List.fold f xs = List.fold f (remove1 x ys)" by (auto intro: Cons.hyps) haftmann@45989: ultimately show ?case by simp haftmann@45989: qed haftmann@45989: haftmann@39533: context linorder haftmann@39533: begin haftmann@39533: haftmann@40210: lemma multiset_of_insort [simp]: haftmann@39533: "multiset_of (insort_key k x xs) = {#x#} + multiset_of xs" haftmann@37107: by (induct xs) (simp_all add: ac_simps) haftmann@39533: haftmann@40210: lemma multiset_of_sort [simp]: haftmann@39533: "multiset_of (sort_key k xs) = multiset_of xs" haftmann@37107: by (induct xs) (simp_all add: ac_simps) haftmann@37107: haftmann@34943: text {* haftmann@34943: This lemma shows which properties suffice to show that a function haftmann@34943: @{text "f"} with @{text "f xs = ys"} behaves like sort. haftmann@34943: *} haftmann@37074: haftmann@39533: lemma properties_for_sort_key: haftmann@39533: assumes "multiset_of ys = multiset_of xs" haftmann@40305: and "\k. k \ set ys \ filter (\x. f k = f x) ys = filter (\x. f k = f x) xs" haftmann@39533: and "sorted (map f ys)" haftmann@39533: shows "sort_key f xs = ys" wenzelm@46921: 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: haftmann@39533: assumes multiset: "multiset_of ys = multiset_of xs" haftmann@39533: and "sorted ys" haftmann@39533: shows "sort xs = ys" haftmann@39533: proof (rule properties_for_sort_key) haftmann@39533: from multiset show "multiset_of ys = multiset_of xs" . haftmann@39533: from `sorted ys` show "sorted (map (\x. x) ys)" by simp haftmann@39533: from multiset have "\k. length (filter (\y. k = y) ys) = length (filter (\x. k = x) xs)" haftmann@39533: by (rule multiset_of_eq_length_filter) haftmann@39533: then have "\k. replicate (length (filter (\y. k = y) ys)) k = replicate (length (filter (\x. k = x) xs)) k" haftmann@39533: by simp haftmann@40305: then show "\k. k \ set ys \ filter (\y. k = y) ys = filter (\x. k = x) xs" haftmann@39533: by (simp add: replicate_length_filter) haftmann@39533: qed haftmann@39533: 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) haftmann@40303: show "multiset_of ?rhs = multiset_of ?lhs" haftmann@40303: by (rule multiset_eqI) (auto simp add: multiset_of_filter) haftmann@40303: next 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: haftmann@40347: 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]: haftmann@40347: "sort_key f xs = (case xs of [] \ [] haftmann@40347: | [x] \ xs haftmann@40347: | [x, y] \ (if f x \ f y then xs else [y, x]) haftmann@40347: | _ \ (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: 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 haftmann@40347: 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: haftmann@35268: lemma multiset_of_remdups_le: "multiset_of (remdups xs) \ multiset_of xs" haftmann@35268: by (induct xs) (auto intro: order_trans) haftmann@34943: haftmann@34943: lemma multiset_of_update: haftmann@34943: "i < length ls \ multiset_of (ls[i := v]) = multiset_of 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@34943: apply (subst add_assoc) haftmann@34943: apply (subst add_commute [of "{#v#}" "{#x#}"]) haftmann@34943: apply (subst add_assoc [symmetric]) haftmann@34943: apply simp haftmann@34943: apply (rule mset_le_multiset_union_diff_commute) haftmann@34943: apply (simp add: mset_le_single nth_mem_multiset_of) haftmann@34943: done haftmann@34943: qed haftmann@34943: qed haftmann@34943: haftmann@34943: lemma multiset_of_swap: haftmann@34943: "i < length ls \ j < length ls \ haftmann@34943: multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls" haftmann@34943: by (cases "i = j") (simp_all add: multiset_of_update nth_mem_multiset_of) haftmann@34943: haftmann@34943: bulwahn@46168: subsubsection {* Association lists -- including code generation *} bulwahn@46168: bulwahn@46168: text {* Preliminaries *} bulwahn@46168: bulwahn@46168: text {* Raw operations on lists *} bulwahn@46168: bulwahn@46168: definition join_raw :: "('key \ 'val \ 'val \ 'val) \ ('key \ 'val) list \ ('key \ 'val) list \ ('key \ 'val) list" bulwahn@46168: where bulwahn@46168: "join_raw f xs ys = foldr (\(k, v). map_default k v (%v'. f k (v', v))) ys xs" bulwahn@46168: bulwahn@46168: lemma join_raw_Nil [simp]: bulwahn@46168: "join_raw f xs [] = xs" bulwahn@46168: by (simp add: join_raw_def) bulwahn@46168: bulwahn@46168: lemma join_raw_Cons [simp]: bulwahn@46168: "join_raw f xs ((k, v) # ys) = map_default k v (%v'. f k (v', v)) (join_raw f xs ys)" bulwahn@46168: by (simp add: join_raw_def) bulwahn@46168: bulwahn@46168: lemma map_of_join_raw: bulwahn@46168: assumes "distinct (map fst ys)" bulwahn@47429: shows "map_of (join_raw f xs ys) x = (case map_of xs x of None => map_of ys x | Some v => bulwahn@47429: (case map_of ys x of None => Some v | Some v' => Some (f x (v, v'))))" bulwahn@46168: using assms bulwahn@46168: apply (induct ys) bulwahn@46168: apply (auto simp add: map_of_map_default split: option.split) bulwahn@46168: apply (metis map_of_eq_None_iff option.simps(2) weak_map_of_SomeI) bulwahn@46168: by (metis Some_eq_map_of_iff map_of_eq_None_iff option.simps(2)) bulwahn@46168: bulwahn@46168: lemma distinct_join_raw: bulwahn@46168: assumes "distinct (map fst xs)" bulwahn@46168: shows "distinct (map fst (join_raw f xs ys))" bulwahn@46168: using assms bulwahn@46168: proof (induct ys) bulwahn@46168: case (Cons y ys) bulwahn@46168: thus ?case by (cases y) (simp add: distinct_map_default) bulwahn@46168: qed auto bulwahn@46168: bulwahn@46168: definition bulwahn@46238: "subtract_entries_raw xs ys = foldr (%(k, v). AList.map_entry k (%v'. v' - v)) ys xs" bulwahn@46168: bulwahn@46168: lemma map_of_subtract_entries_raw: bulwahn@47429: assumes "distinct (map fst ys)" bulwahn@47429: shows "map_of (subtract_entries_raw xs ys) x = (case map_of xs x of None => None | Some v => bulwahn@47429: (case map_of ys x of None => Some v | Some v' => Some (v - v')))" bulwahn@47429: using assms unfolding subtract_entries_raw_def bulwahn@46168: apply (induct ys) bulwahn@46168: apply auto bulwahn@46168: apply (simp split: option.split) bulwahn@46168: apply (simp add: map_of_map_entry) bulwahn@46168: apply (auto split: option.split) bulwahn@46168: apply (metis map_of_eq_None_iff option.simps(3) option.simps(4)) bulwahn@46168: by (metis map_of_eq_None_iff option.simps(4) option.simps(5)) bulwahn@46168: bulwahn@46168: lemma distinct_subtract_entries_raw: bulwahn@46168: assumes "distinct (map fst xs)" bulwahn@46168: shows "distinct (map fst (subtract_entries_raw xs ys))" bulwahn@46168: using assms bulwahn@46168: unfolding subtract_entries_raw_def by (induct ys) (auto simp add: distinct_map_entry) bulwahn@46168: bulwahn@47179: text {* Operations on alists with distinct keys *} bulwahn@46168: kuncar@47308: lift_definition join :: "('a \ 'b \ 'b \ 'b) \ ('a, 'b) alist \ ('a, 'b) alist \ ('a, 'b) alist" kuncar@47308: is join_raw bulwahn@47179: by (simp add: distinct_join_raw) bulwahn@46168: kuncar@47308: lift_definition subtract_entries :: "('a, ('b :: minus)) alist \ ('a, 'b) alist \ ('a, 'b) alist" kuncar@47308: is subtract_entries_raw bulwahn@47179: by (simp add: distinct_subtract_entries_raw) bulwahn@46168: bulwahn@46168: text {* Implementing multisets by means of association lists *} haftmann@34943: haftmann@34943: definition count_of :: "('a \ nat) list \ 'a \ nat" where haftmann@34943: "count_of xs x = (case map_of xs x of None \ 0 | Some n \ n)" haftmann@34943: haftmann@34943: lemma count_of_multiset: haftmann@34943: "count_of xs \ multiset" haftmann@34943: proof - haftmann@34943: let ?A = "{x::'a. 0 < (case map_of xs x of None \ 0\nat | Some (n\nat) \ n)}" haftmann@34943: have "?A \ dom (map_of xs)" haftmann@34943: proof haftmann@34943: fix x haftmann@34943: assume "x \ ?A" haftmann@34943: then have "0 < (case map_of xs x of None \ 0\nat | Some (n\nat) \ n)" by simp haftmann@34943: then have "map_of xs x \ None" by (cases "map_of xs x") auto haftmann@34943: then show "x \ dom (map_of xs)" by auto haftmann@34943: qed haftmann@34943: with finite_dom_map_of [of xs] have "finite ?A" haftmann@34943: by (auto intro: finite_subset) haftmann@34943: then show ?thesis nipkow@39302: by (simp add: count_of_def fun_eq_iff multiset_def) haftmann@34943: qed haftmann@34943: haftmann@34943: lemma count_simps [simp]: haftmann@34943: "count_of [] = (\_. 0)" haftmann@34943: "count_of ((x, n) # xs) = (\y. if x = y then n else count_of xs y)" nipkow@39302: by (simp_all add: count_of_def fun_eq_iff) haftmann@34943: haftmann@34943: lemma count_of_empty: haftmann@34943: "x \ fst ` set xs \ count_of xs x = 0" haftmann@34943: by (induct xs) (simp_all add: count_of_def) haftmann@34943: haftmann@34943: lemma count_of_filter: bulwahn@46168: "count_of (List.filter (P \ fst) xs) x = (if P x then count_of xs x else 0)" haftmann@34943: by (induct xs) auto haftmann@34943: bulwahn@46168: lemma count_of_map_default [simp]: bulwahn@46168: "count_of (map_default x b (%x. x + b) xs) y = (if x = y then count_of xs x + b else count_of xs y)" bulwahn@46168: unfolding count_of_def by (simp add: map_of_map_default split: option.split) bulwahn@46168: bulwahn@46168: lemma count_of_join_raw: bulwahn@46168: "distinct (map fst ys) ==> count_of xs x + count_of ys x = count_of (join_raw (%x (x, y). x + y) xs ys) x" bulwahn@46168: unfolding count_of_def by (simp add: map_of_join_raw split: option.split) bulwahn@46168: bulwahn@46168: lemma count_of_subtract_entries_raw: bulwahn@46168: "distinct (map fst ys) ==> count_of xs x - count_of ys x = count_of (subtract_entries_raw xs ys) x" bulwahn@46168: unfolding count_of_def by (simp add: map_of_subtract_entries_raw split: option.split) bulwahn@46168: bulwahn@46168: text {* Code equations for multiset operations *} bulwahn@46168: bulwahn@46168: definition Bag :: "('a, nat) alist \ 'a multiset" where bulwahn@46237: "Bag xs = Abs_multiset (count_of (DAList.impl_of xs))" haftmann@34943: haftmann@34943: code_datatype Bag haftmann@34943: haftmann@34943: lemma count_Bag [simp, code]: bulwahn@46237: "count (Bag xs) = count_of (DAList.impl_of xs)" haftmann@34943: by (simp add: Bag_def count_of_multiset Abs_multiset_inverse) haftmann@34943: haftmann@34943: lemma Mempty_Bag [code]: bulwahn@46394: "{#} = Bag (DAList.empty)" bulwahn@46394: by (simp add: multiset_eq_iff alist.Alist_inverse DAList.empty_def) bulwahn@47143: haftmann@34943: lemma single_Bag [code]: bulwahn@46394: "{#x#} = Bag (DAList.update x 1 DAList.empty)" kuncar@47198: by (simp add: multiset_eq_iff alist.Alist_inverse update.rep_eq empty.rep_eq) bulwahn@46168: bulwahn@46168: lemma union_Bag [code]: bulwahn@46168: "Bag xs + Bag ys = Bag (join (\x (n1, n2). n1 + n2) xs ys)" bulwahn@46168: by (rule multiset_eqI) (simp add: count_of_join_raw alist.Alist_inverse distinct_join_raw join_def) bulwahn@46168: bulwahn@46168: lemma minus_Bag [code]: bulwahn@46168: "Bag xs - Bag ys = Bag (subtract_entries xs ys)" bulwahn@46168: by (rule multiset_eqI) bulwahn@46168: (simp add: count_of_subtract_entries_raw alist.Alist_inverse distinct_subtract_entries_raw subtract_entries_def) haftmann@34943: haftmann@41069: lemma filter_Bag [code]: bulwahn@46237: "Multiset.filter P (Bag xs) = Bag (DAList.filter (P \ fst) xs)" bulwahn@47429: by (rule multiset_eqI) (simp add: count_of_filter DAList.filter.rep_eq) haftmann@34943: haftmann@34943: lemma mset_less_eq_Bag [code]: bulwahn@46237: "Bag xs \ A \ (\(x, n) \ set (DAList.impl_of xs). count_of (DAList.impl_of xs) x \ count A x)" haftmann@34943: (is "?lhs \ ?rhs") haftmann@34943: proof haftmann@34943: assume ?lhs then show ?rhs wenzelm@46730: by (auto simp add: mset_le_def) haftmann@34943: next haftmann@34943: assume ?rhs haftmann@34943: show ?lhs haftmann@34943: proof (rule mset_less_eqI) haftmann@34943: fix x bulwahn@46237: from `?rhs` have "count_of (DAList.impl_of xs) x \ count A x" bulwahn@46237: by (cases "x \ fst ` set (DAList.impl_of xs)") (auto simp add: count_of_empty) haftmann@34943: then show "count (Bag xs) x \ count A x" wenzelm@46730: by (simp add: mset_le_def) haftmann@34943: qed haftmann@34943: qed haftmann@34943: haftmann@38857: instantiation multiset :: (equal) equal haftmann@34943: begin haftmann@34943: haftmann@34943: definition bulwahn@45866: [code]: "HOL.equal A B \ (A::'a multiset) \ B \ B \ A" haftmann@34943: wenzelm@46921: instance wenzelm@46921: by default (simp add: equal_multiset_def eq_iff) haftmann@34943: haftmann@34943: end haftmann@34943: bulwahn@46168: text {* Quickcheck generators *} haftmann@38857: haftmann@34943: definition (in term_syntax) bulwahn@46168: bagify :: "('a\typerep, nat) alist \ (unit \ Code_Evaluation.term) haftmann@34943: \ 'a multiset \ (unit \ Code_Evaluation.term)" where haftmann@34943: [code_unfold]: "bagify xs = Code_Evaluation.valtermify Bag {\} xs" haftmann@34943: haftmann@37751: notation fcomp (infixl "\>" 60) haftmann@37751: notation scomp (infixl "\\" 60) haftmann@34943: haftmann@34943: instantiation multiset :: (random) random haftmann@34943: begin haftmann@34943: haftmann@34943: definition haftmann@37751: "Quickcheck.random i = Quickcheck.random i \\ (\xs. Pair (bagify xs))" haftmann@34943: haftmann@34943: instance .. haftmann@34943: haftmann@34943: end haftmann@34943: haftmann@37751: no_notation fcomp (infixl "\>" 60) haftmann@37751: no_notation scomp (infixl "\\" 60) haftmann@34943: bulwahn@46168: instantiation multiset :: (exhaustive) exhaustive bulwahn@46168: begin bulwahn@46168: bulwahn@46168: definition exhaustive_multiset :: "('a multiset => (bool * term list) option) => code_numeral => (bool * term list) option" bulwahn@46168: where bulwahn@46168: "exhaustive_multiset f i = Quickcheck_Exhaustive.exhaustive (%xs. f (Bag xs)) i" bulwahn@46168: bulwahn@46168: instance .. bulwahn@46168: bulwahn@46168: end bulwahn@46168: bulwahn@46168: instantiation multiset :: (full_exhaustive) full_exhaustive bulwahn@46168: begin bulwahn@46168: bulwahn@46168: definition full_exhaustive_multiset :: "('a multiset * (unit => term) => (bool * term list) option) => code_numeral => (bool * term list) option" bulwahn@46168: where bulwahn@46168: "full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (%xs. f (bagify xs)) i" bulwahn@46168: bulwahn@46168: instance .. bulwahn@46168: bulwahn@46168: end bulwahn@46168: wenzelm@36176: hide_const (open) bagify haftmann@34943: haftmann@34943: haftmann@34943: subsection {* The multiset order *} wenzelm@10249: wenzelm@10249: subsubsection {* Well-foundedness *} wenzelm@10249: haftmann@28708: 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 \ berghofe@23751: (\b. b :# K --> (b, a) \ r)}" wenzelm@10249: haftmann@28708: 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: berghofe@23751: lemma less_add: "(N, M0 + {#a#}) \ mult1 r ==> berghofe@23751: (\M. (M, M0) \ mult1 r \ N = M + {#a#}) \ berghofe@23751: (\K. (\b. b :# K --> (b, a) \ r) \ N = M0 + K)" wenzelm@19582: (is "_ \ ?case1 (mult1 r) \ ?case2") wenzelm@10249: proof (unfold mult1_def) berghofe@23751: let ?r = "\K a. \b. b :# K --> (b, a) \ r" nipkow@11464: let ?R = "\N M. \a M0 K. M = M0 + {#a#} \ N = M0 + K \ ?r K a" berghofe@23751: let ?case1 = "?case1 {(N, M). ?R N M}" wenzelm@10249: berghofe@23751: assume "(N, M0 + {#a#}) \ {(N, M). ?R N M}" wenzelm@18258: then have "\a' M0' K. nipkow@11464: M0 + {#a#} = M0' + {#a'#} \ N = M0' + K \ ?r K a'" by simp wenzelm@18258: then show "?case1 \ ?case2" wenzelm@10249: proof (elim exE conjE) wenzelm@10249: fix a' M0' K wenzelm@10249: assume N: "N = M0' + K" and r: "?r K a'" wenzelm@10249: assume "M0 + {#a#} = M0' + {#a'#}" wenzelm@18258: then have "M0 = M0' \ a = a' \ nipkow@11464: (\K'. M0 = K' + {#a'#} \ M0' = K' + {#a#})" wenzelm@10249: by (simp only: add_eq_conv_ex) wenzelm@18258: then show ?thesis wenzelm@10249: proof (elim disjE conjE exE) wenzelm@10249: assume "M0 = M0'" "a = a'" nipkow@11464: with N r have "?r K a \ N = M0 + K" by simp wenzelm@18258: then have ?case2 .. then show ?thesis .. wenzelm@10249: next wenzelm@10249: fix K' wenzelm@10249: assume "M0' = K' + {#a#}" haftmann@34943: with N have n: "N = K' + K + {#a#}" by (simp add: add_ac) wenzelm@10249: wenzelm@10249: assume "M0 = K' + {#a'#}" wenzelm@10249: with r have "?R (K' + K) M0" by blast wenzelm@18258: with n have ?case1 by simp then show ?thesis .. wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: berghofe@23751: lemma all_accessible: "wf r ==> \M. M \ acc (mult1 r)" wenzelm@10249: proof wenzelm@10249: let ?R = "mult1 r" wenzelm@10249: let ?W = "acc ?R" wenzelm@10249: { wenzelm@10249: fix M M0 a berghofe@23751: assume M0: "M0 \ ?W" berghofe@23751: and wf_hyp: "!!b. (b, a) \ r ==> (\M \ ?W. M + {#b#} \ ?W)" berghofe@23751: and acc_hyp: "\M. (M, M0) \ ?R --> M + {#a#} \ ?W" berghofe@23751: have "M0 + {#a#} \ ?W" berghofe@23751: proof (rule accI [of "M0 + {#a#}"]) wenzelm@10249: fix N berghofe@23751: assume "(N, M0 + {#a#}) \ ?R" berghofe@23751: then have "((\M. (M, M0) \ ?R \ N = M + {#a#}) \ berghofe@23751: (\K. (\b. b :# K --> (b, a) \ r) \ N = M0 + K))" wenzelm@10249: by (rule less_add) berghofe@23751: then show "N \ ?W" wenzelm@10249: proof (elim exE disjE conjE) berghofe@23751: fix M assume "(M, M0) \ ?R" and N: "N = M + {#a#}" berghofe@23751: from acc_hyp have "(M, M0) \ ?R --> M + {#a#} \ ?W" .. berghofe@23751: from this and `(M, M0) \ ?R` have "M + {#a#} \ ?W" .. berghofe@23751: then show "N \ ?W" by (simp only: N) wenzelm@10249: next wenzelm@10249: fix K wenzelm@10249: assume N: "N = M0 + K" berghofe@23751: assume "\b. b :# K --> (b, a) \ r" berghofe@23751: then have "M0 + K \ ?W" wenzelm@10249: proof (induct K) wenzelm@18730: case empty berghofe@23751: from M0 show "M0 + {#} \ ?W" by simp wenzelm@18730: next wenzelm@18730: case (add K x) berghofe@23751: from add.prems have "(x, a) \ r" by simp berghofe@23751: with wf_hyp have "\M \ ?W. M + {#x#} \ ?W" by blast berghofe@23751: moreover from add have "M0 + K \ ?W" by simp berghofe@23751: ultimately have "(M0 + K) + {#x#} \ ?W" .. haftmann@34943: then show "M0 + (K + {#x#}) \ ?W" by (simp only: add_assoc) wenzelm@10249: qed berghofe@23751: then show "N \ ?W" by (simp only: N) wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: } note tedious_reasoning = this wenzelm@10249: berghofe@23751: assume wf: "wf r" wenzelm@10249: fix M berghofe@23751: show "M \ ?W" wenzelm@10249: proof (induct M) berghofe@23751: show "{#} \ ?W" wenzelm@10249: proof (rule accI) berghofe@23751: fix b assume "(b, {#}) \ ?R" berghofe@23751: with not_less_empty show "b \ ?W" by contradiction wenzelm@10249: qed wenzelm@10249: berghofe@23751: fix M a assume "M \ ?W" berghofe@23751: from wf have "\M \ ?W. M + {#a#} \ ?W" wenzelm@10249: proof induct wenzelm@10249: fix a berghofe@23751: assume r: "!!b. (b, a) \ r ==> (\M \ ?W. M + {#b#} \ ?W)" berghofe@23751: show "\M \ ?W. M + {#a#} \ ?W" wenzelm@10249: proof berghofe@23751: fix M assume "M \ ?W" berghofe@23751: then show "M + {#a#} \ ?W" wenzelm@23373: by (rule acc_induct) (rule tedious_reasoning [OF _ r]) wenzelm@10249: qed wenzelm@10249: qed berghofe@23751: from this and `M \ ?W` show "M + {#a#} \ ?W" .. wenzelm@10249: qed wenzelm@10249: qed wenzelm@10249: berghofe@23751: theorem wf_mult1: "wf r ==> wf (mult1 r)" nipkow@26178: by (rule acc_wfI) (rule all_accessible) wenzelm@10249: berghofe@23751: theorem wf_mult: "wf r ==> wf (mult r)" nipkow@26178: unfolding mult_def by (rule wf_trancl) (rule wf_mult1) wenzelm@10249: wenzelm@10249: wenzelm@10249: subsubsection {* Closure-free presentation *} wenzelm@10249: wenzelm@10249: text {* One direction. *} wenzelm@10249: wenzelm@10249: lemma mult_implies_one_step: berghofe@23751: "trans r ==> (M, N) \ mult r ==> nipkow@11464: \I J K. N = I + J \ M = I + K \ J \ {#} \ berghofe@23751: (\k \ set_of K. \j \ set_of J. (k, j) \ r)" nipkow@26178: apply (unfold mult_def mult1_def set_of_def) nipkow@26178: apply (erule converse_trancl_induct, clarify) nipkow@26178: apply (rule_tac x = M0 in exI, simp, clarify) nipkow@26178: apply (case_tac "a :# K") nipkow@26178: apply (rule_tac x = I in exI) nipkow@26178: apply (simp (no_asm)) nipkow@26178: apply (rule_tac x = "(K - {#a#}) + Ka" in exI) haftmann@34943: apply (simp (no_asm_simp) add: add_assoc [symmetric]) nipkow@26178: apply (drule_tac f = "\M. M - {#a#}" in arg_cong) nipkow@26178: apply (simp add: diff_union_single_conv) nipkow@26178: apply (simp (no_asm_use) add: trans_def) nipkow@26178: apply blast nipkow@26178: apply (subgoal_tac "a :# I") nipkow@26178: apply (rule_tac x = "I - {#a#}" in exI) nipkow@26178: apply (rule_tac x = "J + {#a#}" in exI) nipkow@26178: apply (rule_tac x = "K + Ka" in exI) nipkow@26178: apply (rule conjI) nipkow@39302: apply (simp add: multiset_eq_iff split: nat_diff_split) nipkow@26178: apply (rule conjI) nipkow@26178: apply (drule_tac f = "\M. M - {#a#}" 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 nipkow@26178: apply (subgoal_tac "a :# (M0 + {#a#})") nipkow@26178: apply simp nipkow@26178: apply (simp (no_asm)) nipkow@26178: done wenzelm@10249: wenzelm@10249: lemma one_step_implies_mult_aux: berghofe@23751: "trans r ==> berghofe@23751: \I J K. (size J = n \ J \ {#} \ (\k \ set_of K. \j \ set_of J. (k, j) \ r)) berghofe@23751: --> (I + K, I + J) \ mult r" nipkow@26178: apply (induct_tac n, auto) nipkow@26178: apply (frule size_eq_Suc_imp_eq_union, clarify) nipkow@26178: apply (rename_tac "J'", simp) nipkow@26178: apply (erule notE, auto) nipkow@26178: apply (case_tac "J' = {#}") nipkow@26178: apply (simp add: mult_def) nipkow@26178: apply (rule r_into_trancl) nipkow@26178: apply (simp add: mult1_def set_of_def, blast) nipkow@26178: txt {* Now we know @{term "J' \ {#}"}. *} nipkow@26178: apply (cut_tac M = K and P = "\x. (x, a) \ r" in multiset_partition) nipkow@26178: apply (erule_tac P = "\k \ set_of K. ?P k" in rev_mp) nipkow@26178: apply (erule ssubst) nipkow@26178: apply (simp add: Ball_def, auto) nipkow@26178: apply (subgoal_tac nipkow@26178: "((I + {# x :# K. (x, a) \ r #}) + {# x :# K. (x, a) \ r #}, nipkow@26178: (I + {# x :# K. (x, a) \ r #}) + J') \ mult r") nipkow@26178: prefer 2 nipkow@26178: apply force haftmann@34943: 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@26178: apply (simp add: mult1_def set_of_def) nipkow@26178: apply (rule_tac x = a in exI) nipkow@26178: apply (rule_tac x = "I + J'" in exI) haftmann@34943: apply (simp add: add_ac) nipkow@26178: done wenzelm@10249: wenzelm@17161: lemma one_step_implies_mult: berghofe@23751: "trans r ==> J \ {#} ==> \k \ set_of K. \j \ set_of J. (k, j) \ r berghofe@23751: ==> (I + K, I + J) \ mult r" nipkow@26178: using one_step_implies_mult_aux by blast wenzelm@10249: wenzelm@10249: wenzelm@10249: subsubsection {* Partial-order properties *} wenzelm@10249: haftmann@35273: definition less_multiset :: "'a\order multiset \ 'a multiset \ bool" (infix "<#" 50) where haftmann@35273: "M' <# M \ (M', M) \ mult {(x', x). x' < x}" wenzelm@10249: haftmann@35273: definition le_multiset :: "'a\order multiset \ 'a multiset \ bool" (infix "<=#" 50) where haftmann@35273: "M' <=# M \ M' <# M \ M' = M" haftmann@35273: haftmann@35308: notation (xsymbols) less_multiset (infix "\#" 50) haftmann@35308: notation (xsymbols) le_multiset (infix "\#" 50) wenzelm@10249: haftmann@35268: interpretation multiset_order: order le_multiset less_multiset haftmann@35268: proof - haftmann@35268: have irrefl: "\M :: 'a multiset. \ M \# M" haftmann@35268: proof haftmann@35268: fix M :: "'a multiset" haftmann@35268: 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 haftmann@35268: \ J \ {#} \ (\k\set_of K. \j\set_of 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" haftmann@35268: and "J \ {#}" and "(\k\set_of K. \j\set_of J. (k, j) \ {(x, y). x < y})" by blast haftmann@35268: then have aux1: "K \ {#}" and aux2: "\k\set_of K. \j\set_of K. k < j" by auto haftmann@35268: have "finite (set_of K)" by simp haftmann@35268: moreover note aux2 haftmann@35268: ultimately have "set_of K = {}" haftmann@35268: by (induct rule: finite_induct) (auto intro: order_less_trans) haftmann@35268: with aux1 show False by simp haftmann@35268: qed haftmann@35268: have trans: "\K M N :: 'a multiset. K \# M \ M \# N \ K \# N" 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@46921: by default (auto simp add: le_multiset_def irrefl dest: trans) haftmann@35268: qed wenzelm@10249: wenzelm@46730: lemma mult_less_irrefl [elim!]: "M \# (M::'a::order multiset) ==> R" wenzelm@46730: by simp haftmann@26567: wenzelm@10249: wenzelm@10249: subsubsection {* Monotonicity of multiset union *} wenzelm@10249: wenzelm@46730: 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@34943: apply (simp add: add_assoc) nipkow@26178: done wenzelm@10249: haftmann@35268: 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: haftmann@35268: lemma union_less_mono1: "B \# D ==> B + C \# D + (C::'a::order multiset)" haftmann@34943: apply (subst add_commute [of B C]) haftmann@34943: 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: haftmann@35268: "A \# C ==> B \# D ==> A + B \# C + (D::'a::order multiset)" 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 haftmann@35268: proof haftmann@35268: qed (auto simp add: le_multiset_def intro: union_less_mono2) wenzelm@26145: paulson@15072: krauss@29125: subsection {* 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: krauss@29125: "(set_of A, set_of B) \ max_strict \ (Z + A, Z + B) \ ms_strict" krauss@29125: unfolding ms_strict_def krauss@29125: by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases) krauss@29125: krauss@29125: lemma wmsI: krauss@29125: "(set_of A, set_of B) \ max_strict \ A = {#} \ B = {#} krauss@29125: \ (Z + A, Z + B) \ ms_weak" krauss@29125: unfolding ms_weak_def ms_strict_def krauss@29125: by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult) krauss@29125: krauss@29125: inductive pw_leq krauss@29125: where krauss@29125: pw_leq_empty: "pw_leq {#} {#}" krauss@29125: | pw_leq_step: "\(x,y) \ pair_leq; pw_leq X Y \ \ pw_leq ({#x#} + X) ({#y#} + Y)" krauss@29125: krauss@29125: lemma pw_leq_lstep: krauss@29125: "(x, y) \ pair_leq \ pw_leq {#x#} {#y#}" krauss@29125: by (drule pw_leq_step) (rule pw_leq_empty, simp) krauss@29125: krauss@29125: lemma pw_leq_split: krauss@29125: assumes "pw_leq X Y" krauss@29125: shows "\A B Z. X = A + Z \ Y = B + Z \ ((set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#}))" krauss@29125: using assms krauss@29125: proof (induct) krauss@29125: case pw_leq_empty thus ?case by auto krauss@29125: next krauss@29125: case (pw_leq_step x y X Y) krauss@29125: then obtain A B Z where krauss@29125: [simp]: "X = A + Z" "Y = B + Z" krauss@29125: and 1[simp]: "(set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#})" krauss@29125: by auto krauss@29125: from pw_leq_step have "x = y \ (x, y) \ pair_less" krauss@29125: unfolding pair_leq_def by auto krauss@29125: thus ?case krauss@29125: proof krauss@29125: assume [simp]: "x = y" krauss@29125: have krauss@29125: "{#x#} + X = A + ({#y#}+Z) krauss@29125: \ {#y#} + Y = B + ({#y#}+Z) krauss@29125: \ ((set_of A, set_of B) \ max_strict \ (B = {#} \ A = {#}))" krauss@29125: by (auto simp: add_ac) krauss@29125: thus ?case by (intro exI) krauss@29125: next krauss@29125: assume A: "(x, y) \ pair_less" krauss@29125: let ?A' = "{#x#} + A" and ?B' = "{#y#} + B" krauss@29125: have "{#x#} + X = ?A' + Z" krauss@29125: "{#y#} + Y = ?B' + Z" krauss@29125: by (auto simp add: add_ac) krauss@29125: moreover have krauss@29125: "(set_of ?A', set_of ?B') \ max_strict" krauss@29125: using 1 A unfolding max_strict_def krauss@29125: by (auto elim!: max_ext.cases) krauss@29125: ultimately show ?thesis by blast krauss@29125: qed krauss@29125: qed krauss@29125: krauss@29125: lemma krauss@29125: assumes pwleq: "pw_leq Z Z'" krauss@29125: shows ms_strictI: "(set_of A, set_of B) \ max_strict \ (Z + A, Z' + B) \ ms_strict" krauss@29125: and ms_weakI1: "(set_of A, set_of B) \ max_strict \ (Z + A, Z' + B) \ ms_weak" krauss@29125: and ms_weakI2: "(Z + {#}, Z' + {#}) \ ms_weak" krauss@29125: proof - krauss@29125: from pw_leq_split[OF pwleq] krauss@29125: obtain A' B' Z'' krauss@29125: where [simp]: "Z = A' + Z''" "Z' = B' + Z''" krauss@29125: and mx_or_empty: "(set_of A', set_of B') \ max_strict \ (A' = {#} \ B' = {#})" krauss@29125: by blast krauss@29125: { krauss@29125: assume max: "(set_of A, set_of B) \ max_strict" krauss@29125: from mx_or_empty krauss@29125: have "(Z'' + (A + A'), Z'' + (B + B')) \ ms_strict" krauss@29125: proof krauss@29125: assume max': "(set_of A', set_of B') \ max_strict" krauss@29125: with max have "(set_of (A + A'), set_of (B + B')) \ max_strict" krauss@29125: by (auto simp: max_strict_def intro: max_ext_additive) krauss@29125: thus ?thesis by (rule smsI) krauss@29125: next krauss@29125: assume [simp]: "A' = {#} \ B' = {#}" krauss@29125: show ?thesis by (rule smsI) (auto intro: max) krauss@29125: qed krauss@29125: thus "(Z + A, Z' + B) \ ms_strict" by (simp add:add_ac) krauss@29125: thus "(Z + A, Z' + B) \ ms_weak" by (simp add: ms_weak_def) krauss@29125: } krauss@29125: from mx_or_empty krauss@29125: have "(Z'' + A', Z'' + B') \ ms_weak" by (rule wmsI) krauss@29125: thus "(Z + {#}, Z' + {#}) \ ms_weak" by (simp add:add_ac) krauss@29125: qed krauss@29125: 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: krauss@29125: setup {* krauss@29125: let wenzelm@35402: fun msetT T = Type (@{type_name multiset}, [T]); krauss@29125: wenzelm@35402: fun mk_mset T [] = Const (@{const_abbrev Mempty}, msetT T) krauss@29125: | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x krauss@29125: | mk_mset T (x :: xs) = krauss@29125: Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $ krauss@29125: mk_mset T [x] $ mk_mset T xs krauss@29125: krauss@29125: fun mset_member_tac m i = krauss@29125: (if m <= 0 then krauss@29125: rtac @{thm multi_member_this} i ORELSE rtac @{thm multi_member_last} i krauss@29125: else krauss@29125: rtac @{thm multi_member_skip} i THEN mset_member_tac (m - 1) i) krauss@29125: krauss@29125: val mset_nonempty_tac = krauss@29125: rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single} krauss@29125: krauss@29125: val regroup_munion_conv = wenzelm@35402: Function_Lib.regroup_conv @{const_abbrev Mempty} @{const_name plus} nipkow@39301: (map (fn t => t RS eq_reflection) (@{thms add_ac} @ @{thms empty_neutral})) krauss@29125: krauss@29125: fun unfold_pwleq_tac i = krauss@29125: (rtac @{thm pw_leq_step} i THEN (fn st => unfold_pwleq_tac (i + 1) st)) krauss@29125: ORELSE (rtac @{thm pw_leq_lstep} i) krauss@29125: ORELSE (rtac @{thm pw_leq_empty} i) krauss@29125: krauss@29125: val set_of_simps = [@{thm set_of_empty}, @{thm set_of_single}, @{thm set_of_union}, krauss@29125: @{thm Un_insert_left}, @{thm Un_empty_left}] krauss@29125: in krauss@29125: ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset krauss@29125: { krauss@29125: msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, krauss@29125: mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, krauss@29125: mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_of_simps, wenzelm@30595: smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1}, wenzelm@30595: reduction_pair= @{thm ms_reduction_pair} krauss@29125: }) wenzelm@10249: end krauss@29125: *} krauss@29125: haftmann@34943: haftmann@34943: 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@34943: by (fact add_commute) haftmann@34943: haftmann@34943: lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" haftmann@34943: by (fact add_assoc) haftmann@34943: haftmann@34943: lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" haftmann@34943: 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@34943: by (fact add_imp_eq) haftmann@34943: haftmann@35268: lemma mset_less_trans: "(M::'a multiset) < K \ K < N \ M < N" haftmann@35268: by (fact order_less_trans) haftmann@35268: haftmann@35268: lemma multiset_inter_commute: "A #\ B = B #\ A" haftmann@35268: by (fact inf.commute) haftmann@35268: haftmann@35268: lemma multiset_inter_assoc: "A #\ (B #\ C) = A #\ B #\ C" haftmann@35268: by (fact inf.assoc [symmetric]) haftmann@35268: haftmann@35268: lemma multiset_inter_left_commute: "A #\ (B #\ C) = B #\ (A #\ C)" haftmann@35268: by (fact 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: haftmann@35268: lemma mult_less_not_refl: haftmann@35268: "\ M \# (M::'a::order multiset)" haftmann@35268: by (fact multiset_order.less_irrefl) haftmann@35268: haftmann@35268: lemma mult_less_trans: haftmann@35268: "K \# M ==> M \# N ==> K \# (N::'a::order multiset)" haftmann@35268: by (fact multiset_order.less_trans) haftmann@35268: haftmann@35268: lemma mult_less_not_sym: haftmann@35268: "M \# N ==> \ N \# (M::'a::order multiset)" haftmann@35268: by (fact multiset_order.less_not_sym) haftmann@35268: haftmann@35268: lemma mult_less_asym: haftmann@35268: "M \# N ==> (\ P ==> N \# (M::'a::order multiset)) ==> P" haftmann@35268: by (fact multiset_order.less_asym) haftmann@34943: blanchet@35712: ML {* blanchet@35712: fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) blanchet@35712: (Const _ $ t') = blanchet@35712: let blanchet@35712: val (maybe_opt, ps) = blanchet@35712: Nitpick_Model.dest_plain_fun t' ||> op ~~ blanchet@35712: ||> map (apsnd (snd o HOLogic.dest_number)) blanchet@35712: fun elems_for t = blanchet@35712: case AList.lookup (op =) ps t of blanchet@35712: SOME n => replicate n t blanchet@35712: | NONE => [Const (maybe_name, elem_T --> elem_T) $ t] blanchet@35712: in blanchet@35712: case maps elems_for (all_values elem_T) @ blanchet@37261: (if maybe_opt then [Const (Nitpick_Model.unrep (), elem_T)] blanchet@37261: else []) of blanchet@35712: [] => Const (@{const_name zero_class.zero}, T) blanchet@35712: | ts => foldl1 (fn (t1, t2) => blanchet@35712: Const (@{const_name plus_class.plus}, T --> T --> T) blanchet@35712: $ t1 $ t2) blanchet@35712: (map (curry (op $) (Const (@{const_name single}, blanchet@35712: elem_T --> T))) ts) blanchet@35712: end blanchet@35712: | multiset_postproc _ _ _ _ t = t blanchet@35712: *} blanchet@35712: blanchet@38287: declaration {* blanchet@38287: Nitpick_Model.register_term_postprocessor @{typ "'a multiset"} blanchet@38242: multiset_postproc blanchet@35712: *} blanchet@35712: haftmann@49822: hide_const (open) fold haftmann@49822: blanchet@37169: end haftmann@49388: