Theory HOL-Library.Multiset

(*  Title:      HOL/Library/Multiset.thy
    Author:     Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker
    Author:     Andrei Popescu, TU Muenchen
    Author:     Jasmin Blanchette, Inria, LORIA, MPII
    Author:     Dmitriy Traytel, TU Muenchen
    Author:     Mathias Fleury, MPII
    Author:     Martin Desharnais, MPI-INF Saarbruecken
*)

section ‹(Finite) Multisets›

theory Multiset
  imports Cancellation
begin

subsection ‹The type of multisets›

typedef 'a multiset = {f :: 'a  nat. finite {x. f x > 0}}
  morphisms count Abs_multiset
proof
  show (λx. 0::nat)  {f. finite {x. f x > 0}}
    by simp
qed

setup_lifting type_definition_multiset

lemma count_Abs_multiset:
  count (Abs_multiset f) = f if finite {x. f x > 0}
  by (rule Abs_multiset_inverse) (simp add: that)

lemma multiset_eq_iff: "M = N  (a. count M a = count N a)"
  by (simp only: count_inject [symmetric] fun_eq_iff)

lemma multiset_eqI: "(x. count A x = count B x)  A = B"
  using multiset_eq_iff by auto

text ‹Preservation of the representing set termmultiset.›

lemma diff_preserves_multiset:
  finite {x. 0 < M x - N x} if finite {x. 0 < M x} for M N :: 'a  nat
  using that by (rule rev_finite_subset) auto

lemma filter_preserves_multiset:
  finite {x. 0 < (if P x then M x else 0)} if finite {x. 0 < M x} for M N :: 'a  nat
  using that by (rule rev_finite_subset) auto

lemmas in_multiset = diff_preserves_multiset filter_preserves_multiset


subsection ‹Representing multisets›

text ‹Multiset enumeration›

instantiation multiset :: (type) cancel_comm_monoid_add
begin

lift_definition zero_multiset :: 'a multiset
  is λa. 0
  by simp

abbreviation empty_mset :: 'a multiset ({#})
  where empty_mset  0

lift_definition plus_multiset :: 'a multiset  'a multiset  'a multiset
  is λM N a. M a + N a
  by simp

lift_definition minus_multiset :: 'a multiset  'a multiset  'a multiset
  is λM N a. M a - N a
  by (rule diff_preserves_multiset)

instance
  by (standard; transfer) (simp_all add: fun_eq_iff)

end

context
begin

qualified definition is_empty :: "'a multiset  bool" where
  [code_abbrev]: "is_empty A  A = {#}"

end

lemma add_mset_in_multiset:
  finite {x. 0 < (if x = a then Suc (M x) else M x)}
  if finite {x. 0 < M x}
  using that by (simp add: flip: insert_Collect)

lift_definition add_mset :: "'a  'a multiset  'a multiset" is
  "λa M b. if b = a then Suc (M b) else M b"
by (rule add_mset_in_multiset)

syntax
  "_multiset" :: "args  'a multiset"  ((‹indent=2 notation=‹mixfix multiset enumeration››{#_#}))
translations
  "{#x, xs#}" == "CONST add_mset x {#xs#}"
  "{#x#}" == "CONST add_mset x {#}"

lemma count_empty [simp]: "count {#} a = 0"
  by (simp add: zero_multiset.rep_eq)

lemma count_add_mset [simp]:
  "count (add_mset b A) a = (if b = a then Suc (count A a) else count A a)"
  by (simp add: add_mset.rep_eq)

lemma count_single: "count {#b#} a = (if b = a then 1 else 0)"
  by simp

lemma
  add_mset_not_empty [simp]: add_mset a A  {#} and
  empty_not_add_mset [simp]: "{#}  add_mset a A"
  by (auto simp: multiset_eq_iff)

lemma add_mset_add_mset_same_iff [simp]:
  "add_mset a A = add_mset a B  A = B"
  by (auto simp: multiset_eq_iff)

lemma add_mset_commute:
  "add_mset x (add_mset y M) = add_mset y (add_mset x M)"
  by (auto simp: multiset_eq_iff)


subsection ‹Basic operations›

subsubsection ‹Conversion to set and membership›

definition set_mset :: 'a multiset  'a set
  where set_mset M = {x. count M x > 0}

abbreviation member_mset :: 'a  'a multiset  bool
  where member_mset a M  a  set_mset M

notation
  member_mset  ('(∈#')) and
  member_mset  ((_/ ∈# _) [50, 51] 50)

notation  (ASCII)
  member_mset  ('(:#')) and
  member_mset  ((_/ :# _) [50, 51] 50)

abbreviation not_member_mset :: 'a  'a multiset  bool
  where not_member_mset a M  a  set_mset M

notation
  not_member_mset  ('(∉#')) and
  not_member_mset  ((_/ ∉# _) [50, 51] 50)

notation  (ASCII)
  not_member_mset  ('(~:#')) and
  not_member_mset  ((_/ ~:# _) [50, 51] 50)

context
begin

qualified abbreviation Ball :: "'a multiset  ('a  bool)  bool"
  where "Ball M  Set.Ball (set_mset M)"

qualified abbreviation Bex :: "'a multiset  ('a  bool)  bool"
  where "Bex M  Set.Bex (set_mset M)"

end

syntax
  "_MBall"       :: "pttrn  'a set  bool  bool"      ((3_∈#_./ _) [0, 0, 10] 10)
  "_MBex"        :: "pttrn  'a set  bool  bool"      ((3_∈#_./ _) [0, 0, 10] 10)

syntax  (ASCII)
  "_MBall"       :: "pttrn  'a set  bool  bool"      ((3_:#_./ _) [0, 0, 10] 10)
  "_MBex"        :: "pttrn  'a set  bool  bool"      ((3_:#_./ _) [0, 0, 10] 10)

syntax_consts
  "_MBall"  Multiset.Ball and
  "_MBex"  Multiset.Bex

translations
  "x∈#A. P"  "CONST Multiset.Ball A (λx. P)"
  "x∈#A. P"  "CONST Multiset.Bex A (λx. P)"

print_translation [Syntax_Trans.preserve_binder_abs2_tr' const_syntaxMultiset.Ball syntax_const‹_MBall›,
  Syntax_Trans.preserve_binder_abs2_tr' const_syntaxMultiset.Bex syntax_const‹_MBex›] ― ‹to avoid eta-contraction of body›

lemma count_eq_zero_iff:
  "count M x = 0  x ∉# M"
  by (auto simp add: set_mset_def)

lemma not_in_iff:
  "x ∉# M  count M x = 0"
  by (auto simp add: count_eq_zero_iff)

lemma count_greater_zero_iff [simp]:
  "count M x > 0  x ∈# M"
  by (auto simp add: set_mset_def)

lemma count_inI:
  assumes "count M x = 0  False"
  shows "x ∈# M"
proof (rule ccontr)
  assume "x ∉# M"
  with assms show False by (simp add: not_in_iff)
qed

lemma in_countE:
  assumes "x ∈# M"
  obtains n where "count M x = Suc n"
proof -
  from assms have "count M x > 0" by simp
  then obtain n where "count M x = Suc n"
    using gr0_conv_Suc by blast
  with that show thesis .
qed

lemma count_greater_eq_Suc_zero_iff [simp]:
  "count M x  Suc 0  x ∈# M"
  by (simp add: Suc_le_eq)

lemma count_greater_eq_one_iff [simp]:
  "count M x  1  x ∈# M"
  by simp

lemma set_mset_empty [simp]:
  "set_mset {#} = {}"
  by (simp add: set_mset_def)

lemma set_mset_single:
  "set_mset {#b#} = {b}"
  by (simp add: set_mset_def)

lemma set_mset_eq_empty_iff [simp]:
  "set_mset M = {}  M = {#}"
  by (auto simp add: multiset_eq_iff count_eq_zero_iff)

lemma finite_set_mset [iff]:
  "finite (set_mset M)"
  using count [of M] by simp

lemma set_mset_add_mset_insert [simp]: set_mset (add_mset a A) = insert a (set_mset A)
  by (auto simp flip: count_greater_eq_Suc_zero_iff split: if_splits)

lemma multiset_nonemptyE [elim]:
  assumes "A  {#}"
  obtains x where "x ∈# A"
proof -
  have "x. x ∈# A" by (rule ccontr) (insert assms, auto)
  with that show ?thesis by blast
qed

lemma count_gt_imp_in_mset: "count M x > n  x ∈# M"
  using count_greater_zero_iff by fastforce


subsubsection ‹Union›

lemma count_union [simp]:
  "count (M + N) a = count M a + count N a"
  by (simp add: plus_multiset.rep_eq)

lemma set_mset_union [simp]:
  "set_mset (M + N) = set_mset M  set_mset N"
  by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_union) simp

lemma union_mset_add_mset_left [simp]:
  "add_mset a A + B = add_mset a (A + B)"
  by (auto simp: multiset_eq_iff)

lemma union_mset_add_mset_right [simp]:
  "A + add_mset a B = add_mset a (A + B)"
  by (auto simp: multiset_eq_iff)

(* TODO: reverse arguments to prevent unfolding loop *)
lemma add_mset_add_single: add_mset a A = A + {#a#}
  by (subst union_mset_add_mset_right, subst add.comm_neutral) standard


subsubsection ‹Difference›

instance multiset :: (type) comm_monoid_diff
  by standard (transfer; simp add: fun_eq_iff)

lemma count_diff [simp]:
  "count (M - N) a = count M a - count N a"
  by (simp add: minus_multiset.rep_eq)

lemma add_mset_diff_bothsides:
  add_mset a M - add_mset a A = M - A
  by (auto simp: multiset_eq_iff)

lemma in_diff_count:
  "a ∈# M - N  count N a < count M a"
  by (simp add: set_mset_def)

lemma count_in_diffI:
  assumes "n. count N x = n + count M x  False"
  shows "x ∈# M - N"
proof (rule ccontr)
  assume "x ∉# M - N"
  then have "count N x = (count N x - count M x) + count M x"
    by (simp add: in_diff_count not_less)
  with assms show False by auto
qed

lemma in_diff_countE:
  assumes "x ∈# M - N"
  obtains n where "count M x = Suc n + count N x"
proof -
  from assms have "count M x - count N x > 0" by (simp add: in_diff_count)
  then have "count M x > count N x" by simp
  then obtain n where "count M x = Suc n + count N x"
    using less_iff_Suc_add by auto
  with that show thesis .
qed

lemma in_diffD:
  assumes "a ∈# M - N"
  shows "a ∈# M"
proof -
  have "0  count N a" by simp
  also from assms have "count N a < count M a"
    by (simp add: in_diff_count)
  finally show ?thesis by simp
qed

lemma set_mset_diff:
  "set_mset (M - N) = {a. count N a < count M a}"
  by (simp add: set_mset_def)

lemma diff_empty [simp]: "M - {#} = M  {#} - M = {#}"
  by rule (fact Groups.diff_zero, fact Groups.zero_diff)

lemma diff_cancel: "A - A = {#}"
  by (fact Groups.diff_cancel)

lemma diff_union_cancelR: "M + N - N = (M::'a multiset)"
  by (fact add_diff_cancel_right')

lemma diff_union_cancelL: "N + M - N = (M::'a multiset)"
  by (fact add_diff_cancel_left')

lemma diff_right_commute:
  fixes M N Q :: "'a multiset"
  shows "M - N - Q = M - Q - N"
  by (fact diff_right_commute)

lemma diff_add:
  fixes M N Q :: "'a multiset"
  shows "M - (N + Q) = M - N - Q"
  by (rule sym) (fact diff_diff_add)

lemma insert_DiffM [simp]: "x ∈# M  add_mset x (M - {#x#}) = M"
  by (clarsimp simp: multiset_eq_iff)

lemma insert_DiffM2: "x ∈# M  (M - {#x#}) + {#x#} = M"
  by simp

lemma diff_union_swap: "a  b  add_mset b (M - {#a#}) = add_mset b M - {#a#}"
  by (auto simp add: multiset_eq_iff)

lemma diff_add_mset_swap [simp]: "b ∉# A  add_mset b M - A = add_mset b (M - A)"
  by (auto simp add: multiset_eq_iff simp: not_in_iff)

lemma diff_union_swap2 [simp]: "y ∈# M  add_mset x M - {#y#} = add_mset x (M - {#y#})"
  by (metis add_mset_diff_bothsides diff_union_swap diff_zero insert_DiffM)

lemma diff_diff_add_mset [simp]: "(M::'a multiset) - N - P = M - (N + P)"
  by (rule diff_diff_add)

lemma diff_union_single_conv:
  "a ∈# J  I + J - {#a#} = I + (J - {#a#})"
  by (simp add: multiset_eq_iff Suc_le_eq)

lemma mset_add [elim?]:
  assumes "a ∈# A"
  obtains B where "A = add_mset a B"
proof -
  from assms have "A = add_mset a (A - {#a#})"
    by simp
  with that show thesis .
qed

lemma union_iff:
  "a ∈# A + B  a ∈# A  a ∈# B"
  by auto

lemma count_minus_inter_lt_count_minus_inter_iff:
  "count (M2 - M1) y < count (M1 - M2) y  y ∈# M1 - M2"
  by (meson count_greater_zero_iff gr_implies_not_zero in_diff_count leI order.strict_trans2
      order_less_asym)

lemma minus_inter_eq_minus_inter_iff:
  "(M1 - M2) = (M2 - M1)  set_mset (M1 - M2) = set_mset (M2 - M1)"
  by (metis add.commute count_diff count_eq_zero_iff diff_add_zero in_diff_countE multiset_eq_iff)


subsubsection ‹Min and Max›

abbreviation Min_mset :: "'a::linorder multiset  'a" where
"Min_mset m  Min (set_mset m)"

abbreviation Max_mset :: "'a::linorder multiset  'a" where
"Max_mset m  Max (set_mset m)"

lemma
  Min_in_mset: "M  {#}  Min_mset M ∈# M" and
  Max_in_mset: "M  {#}  Max_mset M ∈# M"
  by simp+


subsubsection ‹Equality of multisets›

lemma single_eq_single [simp]: "{#a#} = {#b#}  a = b"
  by (auto simp add: multiset_eq_iff)

lemma union_eq_empty [iff]: "M + N = {#}  M = {#}  N = {#}"
  by (auto simp add: multiset_eq_iff)

lemma empty_eq_union [iff]: "{#} = M + N  M = {#}  N = {#}"
  by (auto simp add: multiset_eq_iff)

lemma multi_self_add_other_not_self [simp]: "M = add_mset x M  False"
  by (auto simp add: multiset_eq_iff)

lemma add_mset_remove_trivial [simp]: add_mset x M - {#x#} = M
  by (auto simp: multiset_eq_iff)

lemma diff_single_trivial: "¬ x ∈# M  M - {#x#} = M"
  by (auto simp add: multiset_eq_iff not_in_iff)

lemma diff_single_eq_union: "x ∈# M  M - {#x#} = N  M = add_mset x N"
  by auto

lemma union_single_eq_diff: "add_mset x M = N  M = N - {#x#}"
  unfolding add_mset_add_single[of _ M] by (fact add_implies_diff)

lemma union_single_eq_member: "add_mset x M = N  x ∈# N"
  by auto

lemma add_mset_remove_trivial_If:
  "add_mset a (N - {#a#}) = (if a ∈# N then N else add_mset a N)"
  by (simp add: diff_single_trivial)

lemma add_mset_remove_trivial_eq: N = add_mset a (N - {#a#})  a ∈# N
  by (auto simp: add_mset_remove_trivial_If)

lemma union_is_single:
  "M + N = {#a#}  M = {#a#}  N = {#}  M = {#}  N = {#a#}"
  (is "?lhs = ?rhs")
proof
  show ?lhs if ?rhs using that by auto
  show ?rhs if ?lhs
    by (metis Multiset.diff_cancel add.commute add_diff_cancel_left' diff_add_zero diff_single_trivial insert_DiffM that)
qed

lemma single_is_union: "{#a#} = M + N  {#a#} = M  N = {#}  M = {#}  {#a#} = N"
  by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single)

lemma add_eq_conv_diff:
  "add_mset a M = add_mset b N  M = N  a = b  M = add_mset b (N - {#a#})  N = add_mset a (M - {#b#})"
  (is "?lhs  ?rhs")
(* shorter: by (simp add: multiset_eq_iff) fastforce *)
proof
  show ?lhs if ?rhs
    using that
    by (auto simp add: add_mset_commute[of a b])
  show ?rhs if ?lhs
  proof (cases "a = b")
    case True with ?lhs show ?thesis by simp
  next
    case False
    from ?lhs have "a ∈# add_mset b N" by (rule union_single_eq_member)
    with False have "a ∈# N" by auto
    moreover from ?lhs have "M = add_mset b N - {#a#}" by (rule union_single_eq_diff)
    moreover note False
    ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"])
  qed
qed

lemma add_mset_eq_single [iff]: "add_mset b M = {#a#}  b = a  M = {#}"
  by (auto simp: add_eq_conv_diff)

lemma single_eq_add_mset [iff]: "{#a#} = add_mset b M  b = a  M = {#}"
  by (auto simp: add_eq_conv_diff)

lemma insert_noteq_member:
  assumes BC: "add_mset b B = add_mset c C"
   and bnotc: "b  c"
  shows "c ∈# B"
proof -
  have "c ∈# add_mset c C" by simp
  have nc: "¬ c ∈# {#b#}" using bnotc by simp
  then have "c ∈# add_mset b B" using BC by simp
  then show "c ∈# B" using nc by simp
qed

lemma add_eq_conv_ex:
  "(add_mset a M = add_mset b N) =
    (M = N  a = b  (K. M = add_mset b K  N = add_mset a K))"
  by (auto simp add: add_eq_conv_diff)

lemma multi_member_split: "x ∈# M  A. M = add_mset x A"
  by (rule exI [where x = "M - {#x#}"]) simp

lemma multiset_add_sub_el_shuffle:
  assumes "c ∈# B"
    and "b  c"
  shows "add_mset b (B - {#c#}) = add_mset b B - {#c#}"
proof -
  from c ∈# B obtain A where B: "B = add_mset c A"
    by (blast dest: multi_member_split)
  have "add_mset b A = add_mset c (add_mset b A) - {#c#}" by simp
  then have "add_mset b A = add_mset b (add_mset c A) - {#c#}"
    by (simp add: b  c)
  then show ?thesis using B by simp
qed

lemma add_mset_eq_singleton_iff[iff]:
  "add_mset x M = {#y#}  M = {#}  x = y"
  by auto


subsubsection ‹Pointwise ordering induced by count›

definition subseteq_mset :: "'a multiset  'a multiset  bool"  (infix ⊆# 50)
  where "A ⊆# B  (a. count A a  count B a)"

definition subset_mset :: "'a multiset  'a multiset  bool" (infix ⊂# 50)
  where "A ⊂# B  A ⊆# B  A  B"

abbreviation (input) supseteq_mset :: "'a multiset  'a multiset  bool"  (infix ⊇# 50)
  where "supseteq_mset A B  B ⊆# A"

abbreviation (input) supset_mset :: "'a multiset  'a multiset  bool"  (infix ⊃# 50)
  where "supset_mset A B  B ⊂# A"

notation (input)
  subseteq_mset  (infix ≤# 50) and
  supseteq_mset  (infix ≥# 50)

notation (ASCII)
  subseteq_mset  (infix <=# 50) and
  subset_mset  (infix <# 50) and
  supseteq_mset  (infix >=# 50) and
  supset_mset  (infix ># 50)

global_interpretation subset_mset: ordering (⊆#) (⊂#)
  by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order.trans order.antisym)

interpretation subset_mset: ordered_ab_semigroup_add_imp_le (+) (-) (⊆#) (⊂#)
  by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym)
    ― ‹FIXME: avoid junk stemming from type class interpretation›

interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "(+)" 0 "(-)" "(⊆#)" "(⊂#)"
  by standard
    ― ‹FIXME: avoid junk stemming from type class interpretation›

lemma mset_subset_eqI:
  "(a. count A a  count B a)  A ⊆# B"
  by (simp add: subseteq_mset_def)

lemma mset_subset_eq_count:
  "A ⊆# B  count A a  count B a"
  by (simp add: subseteq_mset_def)

lemma mset_subset_eq_exists_conv: "(A::'a multiset) ⊆# B  (C. B = A + C)"
  unfolding subseteq_mset_def
  by (metis add_diff_cancel_left' count_diff count_union le_Suc_ex le_add_same_cancel1 multiset_eq_iff zero_le)

interpretation subset_mset: ordered_cancel_comm_monoid_diff "(+)" 0 "(⊆#)" "(⊂#)" "(-)"
  by standard (simp, fact mset_subset_eq_exists_conv)
    ― ‹FIXME: avoid junk stemming from type class interpretation›

declare subset_mset.add_diff_assoc[simp] subset_mset.add_diff_assoc2[simp]

lemma mset_subset_eq_mono_add_right_cancel: "(A::'a multiset) + C ⊆# B + C  A ⊆# B"
   by (fact subset_mset.add_le_cancel_right)

lemma mset_subset_eq_mono_add_left_cancel: "C + (A::'a multiset) ⊆# C + B  A ⊆# B"
   by (fact subset_mset.add_le_cancel_left)

lemma mset_subset_eq_mono_add: "(A::'a multiset) ⊆# B  C ⊆# D  A + C ⊆# B + D"
   by (fact subset_mset.add_mono)

lemma mset_subset_eq_add_left: "(A::'a multiset) ⊆# A + B"
   by simp

lemma mset_subset_eq_add_right: "B ⊆# (A::'a multiset) + B"
   by simp

lemma single_subset_iff [simp]:
  "{#a#} ⊆# M  a ∈# M"
  by (auto simp add: subseteq_mset_def Suc_le_eq)

lemma mset_subset_eq_single: "a ∈# B  {#a#} ⊆# B"
  by simp

lemma mset_subset_eq_add_mset_cancel: add_mset a A ⊆# add_mset a B  A ⊆# B
  unfolding add_mset_add_single[of _ A] add_mset_add_single[of _ B]
  by (rule mset_subset_eq_mono_add_right_cancel)

lemma multiset_diff_union_assoc:
  fixes A B C D :: "'a multiset"
  shows "C ⊆# B  A + B - C = A + (B - C)"
  by (fact subset_mset.diff_add_assoc)

lemma mset_subset_eq_multiset_union_diff_commute:
  fixes A B C D :: "'a multiset"
  shows "B ⊆# A  A - B + C = A + C - B"
  by (fact subset_mset.add_diff_assoc2)

lemma diff_subset_eq_self[simp]:
  "(M::'a multiset) - N ⊆# M"
  by (simp add: subseteq_mset_def)

lemma mset_subset_eqD:
  assumes "A ⊆# B" and "x ∈# A"
  shows "x ∈# B"
proof -
  from x ∈# A have "count A x > 0" by simp
  also from A ⊆# B have "count A x  count B x"
    by (simp add: subseteq_mset_def)
  finally show ?thesis by simp
qed

lemma mset_subsetD:
  "A ⊂# B  x ∈# A  x ∈# B"
  by (auto intro: mset_subset_eqD [of A])

lemma set_mset_mono:
  "A ⊆# B  set_mset A  set_mset B"
  by (metis mset_subset_eqD subsetI)

lemma mset_subset_eq_insertD:
  assumes "add_mset x A ⊆# B"
  shows "x ∈# B  A ⊂# B"
proof 
  show "x ∈# B"
    using assms by (simp add: mset_subset_eqD)
  have "A ⊆# add_mset x A"
    by (metis (no_types) add_mset_add_single mset_subset_eq_add_left)
  then have "A ⊂# add_mset x A"
    by (meson multi_self_add_other_not_self subset_mset.le_imp_less_or_eq)
  then show "A ⊂# B"
    using assms subset_mset.strict_trans2 by blast
qed

lemma mset_subset_insertD:
  "add_mset x A ⊂# B  x ∈# B  A ⊂# B"
  by (rule mset_subset_eq_insertD) simp

lemma mset_subset_of_empty[simp]: "A ⊂# {#}  False"
  by (simp only: subset_mset.not_less_zero)

lemma empty_subset_add_mset[simp]: "{#} ⊂# add_mset x M"
  by (auto intro: subset_mset.gr_zeroI)

lemma empty_le: "{#} ⊆# A"
  by (fact subset_mset.zero_le)

lemma insert_subset_eq_iff:
  "add_mset a A ⊆# B  a ∈# B  A ⊆# B - {#a#}"
  using mset_subset_eq_insertD subset_mset.le_diff_conv2 by fastforce

lemma insert_union_subset_iff:
  "add_mset a A ⊂# B  a ∈# B  A ⊂# B - {#a#}"
  by (auto simp add: insert_subset_eq_iff subset_mset_def)

lemma subset_eq_diff_conv:
  "A - C ⊆# B  A ⊆# B + C"
  by (simp add: subseteq_mset_def le_diff_conv)

lemma multi_psub_of_add_self [simp]: "A ⊂# add_mset x A"
  by (auto simp: subset_mset_def subseteq_mset_def)

lemma multi_psub_self: "A ⊂# A = False"
  by simp

lemma mset_subset_add_mset [simp]: "add_mset x N ⊂# add_mset x M  N ⊂# M"
  unfolding add_mset_add_single[of _ N] add_mset_add_single[of _ M]
  by (fact subset_mset.add_less_cancel_right)

lemma mset_subset_diff_self: "c ∈# B  B - {#c#} ⊂# B"
  by (auto simp: subset_mset_def elim: mset_add)

lemma Diff_eq_empty_iff_mset: "A - B = {#}  A ⊆# B"
  by (auto simp: multiset_eq_iff subseteq_mset_def)

lemma add_mset_subseteq_single_iff[iff]: "add_mset a M ⊆# {#b#}  M = {#}  a = b"
proof
  assume A: "add_mset a M ⊆# {#b#}"
  then have a = b
    by (auto dest: mset_subset_eq_insertD)
  then show "M={#}  a=b"
    using A by (simp add: mset_subset_eq_add_mset_cancel)
qed simp

lemma nonempty_subseteq_mset_eq_single: "M  {#}  M ⊆# {#x#}  M = {#x#}"
  by (cases M) (metis single_is_union subset_mset.less_eqE)

lemma nonempty_subseteq_mset_iff_single: "(M  {#}  M ⊆# {#x#}  P)  M = {#x#}  P"
  by (cases M) (metis empty_not_add_mset nonempty_subseteq_mset_eq_single subset_mset.order_refl)


subsubsection ‹Intersection and bounded union›

definition inter_mset :: 'a multiset  'a multiset  'a multiset  (infixl ∩# 70)
  where A ∩# B = A - (A - B)

lemma count_inter_mset [simp]:
  count (A ∩# B) x = min (count A x) (count B x)
  by (simp add: inter_mset_def)

(*global_interpretation subset_mset: semilattice_order ‹(∩#)› ‹(⊆#)› ‹(⊂#)›
  by standard (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def min_def)*)

interpretation subset_mset: semilattice_inf (∩#) (⊆#) (⊂#)
  by standard (simp_all add: multiset_eq_iff subseteq_mset_def)
  ― ‹FIXME: avoid junk stemming from type class interpretation›

definition union_mset :: 'a multiset  'a multiset  'a multiset  (infixl ∪# 70)
  where A ∪# B = A + (B - A)

lemma count_union_mset [simp]:
  count (A ∪# B) x = max (count A x) (count B x)
  by (simp add: union_mset_def)

global_interpretation subset_mset: semilattice_neutr_order (∪#) {#} (⊇#) (⊃#)
proof 
  show "a b. (b ⊆# a) = (a = a ∪# b)"
    by (simp add: Diff_eq_empty_iff_mset union_mset_def)
  show "a b. (b ⊂# a) = (a = a ∪# b  a  b)"
    by (metis Diff_eq_empty_iff_mset add_cancel_left_right subset_mset_def union_mset_def)
qed (auto simp: multiset_eqI union_mset_def)

interpretation subset_mset: semilattice_sup (∪#) (⊆#) (⊂#)
proof -
  have [simp]: "m  n  q  n  m + (q - m)  n" for m n q :: nat
    by arith
  show "class.semilattice_sup (∪#) (⊆#) (⊂#)"
    by standard (auto simp add: union_mset_def subseteq_mset_def)
qed ― ‹FIXME: avoid junk stemming from type class interpretation›

interpretation subset_mset: bounded_lattice_bot "(∩#)" "(⊆#)" "(⊂#)"
  "(∪#)" "{#}"
  by standard auto
    ― ‹FIXME: avoid junk stemming from type class interpretation›


subsubsection ‹Additional intersection facts›

lemma set_mset_inter [simp]:
  "set_mset (A ∩# B) = set_mset A  set_mset B"
  by (simp only: set_mset_def) auto

lemma diff_intersect_left_idem [simp]:
  "M - M ∩# N = M - N"
  by (simp add: multiset_eq_iff min_def)

lemma diff_intersect_right_idem [simp]:
  "M - N ∩# M = M - N"
  by (simp add: multiset_eq_iff min_def)

lemma multiset_inter_single[simp]: "a  b  {#a#} ∩# {#b#} = {#}"
  by (rule multiset_eqI) auto

lemma multiset_union_diff_commute:
  assumes "B ∩# C = {#}"
  shows "A + B - C = A - C + B"
proof (rule multiset_eqI)
  fix x
  from assms have "min (count B x) (count C x) = 0"
    by (auto simp add: multiset_eq_iff)
  then have "count B x = 0  count C x = 0"
    unfolding min_def by (auto split: if_splits)
  then show "count (A + B - C) x = count (A - C + B) x"
    by auto
qed

lemma disjunct_not_in:
  "A ∩# B = {#}  (a. a ∉# A  a ∉# B)"
  by (metis disjoint_iff set_mset_eq_empty_iff set_mset_inter)

lemma inter_mset_empty_distrib_right: "A ∩# (B + C) = {#}  A ∩# B = {#}  A ∩# C = {#}"
  by (meson disjunct_not_in union_iff)

lemma inter_mset_empty_distrib_left: "(A + B) ∩# C = {#}  A ∩# C = {#}  B ∩# C = {#}"
  by (meson disjunct_not_in union_iff)

lemma add_mset_inter_add_mset [simp]:
  "add_mset a A ∩# add_mset a B = add_mset a (A ∩# B)"
  by (rule multiset_eqI) simp

lemma add_mset_disjoint [simp]:
  "add_mset a A ∩# B = {#}  a ∉# B  A ∩# B = {#}"
  "{#} = add_mset a A ∩# B  a ∉# B  {#} = A ∩# B"
  by (auto simp: disjunct_not_in)

lemma disjoint_add_mset [simp]:
  "B ∩# add_mset a A = {#}  a ∉# B  B ∩# A = {#}"
  "{#} = A ∩# add_mset b B  b ∉# A  {#} = A ∩# B"
  by (auto simp: disjunct_not_in)

lemma inter_add_left1: "¬ x ∈# N  (add_mset x M) ∩# N = M ∩# N"
  by (simp add: multiset_eq_iff not_in_iff)

lemma inter_add_left2: "x ∈# N  (add_mset x M) ∩# N = add_mset x (M ∩# (N - {#x#}))"
  by (auto simp add: multiset_eq_iff elim: mset_add)

lemma inter_add_right1: "¬ x ∈# N  N ∩# (add_mset x M) = N ∩# M"
  by (simp add: multiset_eq_iff not_in_iff)

lemma inter_add_right2: "x ∈# N  N ∩# (add_mset x M) = add_mset x ((N - {#x#}) ∩# M)"
  by (auto simp add: multiset_eq_iff elim: mset_add)

lemma disjunct_set_mset_diff:
  assumes "M ∩# N = {#}"
  shows "set_mset (M - N) = set_mset M"
proof (rule set_eqI)
  fix a
  from assms have "a ∉# M  a ∉# N"
    by (simp add: disjunct_not_in)
  then show "a ∈# M - N  a ∈# M"
    by (auto dest: in_diffD) (simp add: in_diff_count not_in_iff)
qed

lemma at_most_one_mset_mset_diff:
  assumes "a ∉# M - {#a#}"
  shows "set_mset (M - {#a#}) = set_mset M - {a}"
  using assms by (auto simp add: not_in_iff in_diff_count set_eq_iff)

lemma more_than_one_mset_mset_diff:
  assumes "a ∈# M - {#a#}"
  shows "set_mset (M - {#a#}) = set_mset M"
proof (rule set_eqI)
  fix b
  have "Suc 0 < count M b  count M b > 0" by arith
  then show "b ∈# M - {#a#}  b ∈# M"
    using assms by (auto simp add: in_diff_count)
qed

lemma inter_iff:
  "a ∈# A ∩# B  a ∈# A  a ∈# B"
  by simp

lemma inter_union_distrib_left:
  "A ∩# B + C = (A + C) ∩# (B + C)"
  by (simp add: multiset_eq_iff min_add_distrib_left)

lemma inter_union_distrib_right:
  "C + A ∩# B = (C + A) ∩# (C + B)"
  using inter_union_distrib_left [of A B C] by (simp add: ac_simps)

lemma inter_subset_eq_union:
  "A ∩# B ⊆# A + B"
  by (auto simp add: subseteq_mset_def)


subsubsection ‹Additional bounded union facts›

lemma set_mset_sup [simp]:
  set_mset (A ∪# B) = set_mset A  set_mset B
  by (simp only: set_mset_def) (auto simp add: less_max_iff_disj)

lemma sup_union_left1 [simp]: "¬ x ∈# N  (add_mset x M) ∪# N = add_mset x (M ∪# N)"
  by (simp add: multiset_eq_iff not_in_iff)

lemma sup_union_left2: "x ∈# N  (add_mset x M) ∪# N = add_mset x (M ∪# (N - {#x#}))"
  by (simp add: multiset_eq_iff)

lemma sup_union_right1 [simp]: "¬ x ∈# N  N ∪# (add_mset x M) = add_mset x (N ∪# M)"
  by (simp add: multiset_eq_iff not_in_iff)

lemma sup_union_right2: "x ∈# N  N ∪# (add_mset x M) = add_mset x ((N - {#x#}) ∪# M)"
  by (simp add: multiset_eq_iff)

lemma sup_union_distrib_left:
  "A ∪# B + C = (A + C) ∪# (B + C)"
  by (simp add: multiset_eq_iff max_add_distrib_left)

lemma union_sup_distrib_right:
  "C + A ∪# B = (C + A) ∪# (C + B)"
  using sup_union_distrib_left [of A B C] by (simp add: ac_simps)

lemma union_diff_inter_eq_sup:
  "A + B - A ∩# B = A ∪# B"
  by (auto simp add: multiset_eq_iff)

lemma union_diff_sup_eq_inter:
  "A + B - A ∪# B = A ∩# B"
  by (auto simp add: multiset_eq_iff)

lemma add_mset_union:
  add_mset a A ∪# add_mset a B = add_mset a (A ∪# B)
  by (auto simp: multiset_eq_iff max_def)


subsection ‹Replicate and repeat operations›

definition replicate_mset :: "nat  'a  'a multiset" where
  "replicate_mset n x = (add_mset x ^^ n) {#}"

lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}"
  unfolding replicate_mset_def by simp

lemma replicate_mset_Suc [simp]: "replicate_mset (Suc n) x = add_mset x (replicate_mset n x)"
  unfolding replicate_mset_def by (induct n) (auto intro: add.commute)

lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)"
  unfolding replicate_mset_def by (induct n) auto

lift_definition repeat_mset :: nat  'a multiset  'a multiset
  is λn M a. n * M a by simp

lemma count_repeat_mset [simp]: "count (repeat_mset i A) a = i * count A a"
  by transfer rule

lemma repeat_mset_0 [simp]:
  repeat_mset 0 M = {#}
  by transfer simp

lemma repeat_mset_Suc [simp]:
  repeat_mset (Suc n) M = M + repeat_mset n M
  by transfer simp

lemma repeat_mset_right [simp]: "repeat_mset a (repeat_mset b A) = repeat_mset (a * b) A"
  by (auto simp: multiset_eq_iff left_diff_distrib')

lemma left_diff_repeat_mset_distrib': repeat_mset (i - j) u = repeat_mset i u - repeat_mset j u
  by (auto simp: multiset_eq_iff left_diff_distrib')

lemma left_add_mult_distrib_mset:
  "repeat_mset i u + (repeat_mset j u + k) = repeat_mset (i+j) u + k"
  by (auto simp: multiset_eq_iff add_mult_distrib)

lemma repeat_mset_distrib:
  "repeat_mset (m + n) A = repeat_mset m A + repeat_mset n A"
  by (auto simp: multiset_eq_iff Nat.add_mult_distrib)

lemma repeat_mset_distrib2[simp]:
  "repeat_mset n (A + B) = repeat_mset n A + repeat_mset n B"
  by (auto simp: multiset_eq_iff add_mult_distrib2)

lemma repeat_mset_replicate_mset[simp]:
  "repeat_mset n {#a#} = replicate_mset n a"
  by (auto simp: multiset_eq_iff)

lemma repeat_mset_distrib_add_mset[simp]:
  "repeat_mset n (add_mset a A) = replicate_mset n a + repeat_mset n A"
  by (auto simp: multiset_eq_iff)

lemma repeat_mset_empty[simp]: "repeat_mset n {#} = {#}"
  by transfer simp


subsubsection ‹Simprocs›

lemma repeat_mset_iterate_add: repeat_mset n M = iterate_add n M
  unfolding iterate_add_def by (induction n) auto

lemma mset_subseteq_add_iff1:
  "j  (i::nat)  (repeat_mset i u + m ⊆# repeat_mset j u + n) = (repeat_mset (i-j) u + m ⊆# n)"
  by (auto simp add: subseteq_mset_def nat_le_add_iff1)

lemma mset_subseteq_add_iff2:
  "i  (j::nat)  (repeat_mset i u + m ⊆# repeat_mset j u + n) = (m ⊆# repeat_mset (j-i) u + n)"
  by (auto simp add: subseteq_mset_def nat_le_add_iff2)

lemma mset_subset_add_iff1:
  "j  (i::nat)  (repeat_mset i u + m ⊂# repeat_mset j u + n) = (repeat_mset (i-j) u + m ⊂# n)"
  unfolding subset_mset_def repeat_mset_iterate_add
  by (simp add: iterate_add_eq_add_iff1 mset_subseteq_add_iff1[unfolded repeat_mset_iterate_add])

lemma mset_subset_add_iff2:
  "i  (j::nat)  (repeat_mset i u + m ⊂# repeat_mset j u + n) = (m ⊂# repeat_mset (j-i) u + n)"
  unfolding subset_mset_def repeat_mset_iterate_add
  by (simp add: iterate_add_eq_add_iff2 mset_subseteq_add_iff2[unfolded repeat_mset_iterate_add])

ML_file ‹multiset_simprocs.ML›

lemma add_mset_replicate_mset_safe[cancelation_simproc_pre]: NO_MATCH {#} M  add_mset a M = {#a#} + M
  by simp

declare repeat_mset_iterate_add[cancelation_simproc_pre]

declare iterate_add_distrib[cancelation_simproc_pre]
declare repeat_mset_iterate_add[symmetric, cancelation_simproc_post]

declare add_mset_not_empty[cancelation_simproc_eq_elim]
    empty_not_add_mset[cancelation_simproc_eq_elim]
    subset_mset.le_zero_eq[cancelation_simproc_eq_elim]
    empty_not_add_mset[cancelation_simproc_eq_elim]
    add_mset_not_empty[cancelation_simproc_eq_elim]
    subset_mset.le_zero_eq[cancelation_simproc_eq_elim]
    le_zero_eq[cancelation_simproc_eq_elim]

simproc_setup mseteq_cancel
  ("(l::'a multiset) + m = n" | "(l::'a multiset) = m + n" |
   "add_mset a m = n" | "m = add_mset a n" |
   "replicate_mset p a = n" | "m = replicate_mset p a" |
   "repeat_mset p m = n" | "m = repeat_mset p m") =
  K Cancel_Simprocs.eq_cancel

simproc_setup msetsubset_cancel
  ("(l::'a multiset) + m ⊂# n" | "(l::'a multiset) ⊂# m + n" |
   "add_mset a m ⊂# n" | "m ⊂# add_mset a n" |
   "replicate_mset p r ⊂# n" | "m ⊂# replicate_mset p r" |
   "repeat_mset p m ⊂# n" | "m ⊂# repeat_mset p m") =
  K Multiset_Simprocs.subset_cancel_msets

simproc_setup msetsubset_eq_cancel
  ("(l::'a multiset) + m ⊆# n" | "(l::'a multiset) ⊆# m + n" |
   "add_mset a m ⊆# n" | "m ⊆# add_mset a n" |
   "replicate_mset p r ⊆# n" | "m ⊆# replicate_mset p r" |
   "repeat_mset p m ⊆# n" | "m ⊆# repeat_mset p m") =
  K Multiset_Simprocs.subseteq_cancel_msets

simproc_setup msetdiff_cancel
  ("((l::'a multiset) + m) - n" | "(l::'a multiset) - (m + n)" |
   "add_mset a m - n" | "m - add_mset a n" |
   "replicate_mset p r - n" | "m - replicate_mset p r" |
   "repeat_mset p m - n" | "m - repeat_mset p m") =
  K Cancel_Simprocs.diff_cancel


subsubsection ‹Conditionally complete lattice›

instantiation multiset :: (type) Inf
begin

lift_definition Inf_multiset :: "'a multiset set  'a multiset" is
  "λA i. if A = {} then 0 else Inf ((λf. f i) ` A)"
proof -
  fix A :: "('a  nat) set"
  assume *: "f. f  A  finite {x. 0 < f x}"
  show finite {i. 0 < (if A = {} then 0 else INF fA. f i)}
  proof (cases "A = {}")
    case False
    then obtain f where "f  A" by blast
    hence "{i. Inf ((λf. f i) ` A) > 0}  {i. f i > 0}"
      by (auto intro: less_le_trans[OF _ cInf_lower])
    moreover from f  A * have "finite " by simp
    ultimately have "finite {i. Inf ((λf. f i) ` A) > 0}" by (rule finite_subset)
    with False show ?thesis by simp
  qed simp_all
qed

instance ..

end

lemma Inf_multiset_empty: "Inf {} = {#}"
  by transfer simp_all

lemma count_Inf_multiset_nonempty: "A  {}  count (Inf A) x = Inf ((λX. count X x) ` A)"
  by transfer simp_all


instantiation multiset :: (type) Sup
begin

definition Sup_multiset :: "'a multiset set  'a multiset" where
  "Sup_multiset A = (if A  {}  subset_mset.bdd_above A then
           Abs_multiset (λi. Sup ((λX. count X i) ` A)) else {#})"

lemma Sup_multiset_empty: "Sup {} = {#}"
  by (simp add: Sup_multiset_def)

lemma Sup_multiset_unbounded: "¬ subset_mset.bdd_above A  Sup A = {#}"
  by (simp add: Sup_multiset_def)

instance ..

end

lemma bdd_above_multiset_imp_bdd_above_count:
  assumes "subset_mset.bdd_above (A :: 'a multiset set)"
  shows   "bdd_above ((λX. count X x) ` A)"
proof -
  from assms obtain Y where Y: "XA. X ⊆# Y"
    by (meson subset_mset.bdd_above.E)
  hence "count X x  count Y x" if "X  A" for X
    using that by (auto intro: mset_subset_eq_count)
  thus ?thesis by (intro bdd_aboveI[of _ "count Y x"]) auto
qed

lemma bdd_above_multiset_imp_finite_support:
  assumes "A  {}" "subset_mset.bdd_above (A :: 'a multiset set)"
  shows   "finite (XA. {x. count X x > 0})"
proof -
  from assms obtain Y where Y: "XA. X ⊆# Y"
    by (meson subset_mset.bdd_above.E)
  hence "count X x  count Y x" if "X  A" for X x
    using that by (auto intro: mset_subset_eq_count)
  hence "(XA. {x. count X x > 0})  {x. count Y x > 0}"
    by safe (erule less_le_trans)
  moreover have "finite " by simp
  ultimately show ?thesis by (rule finite_subset)
qed

lemma Sup_multiset_in_multiset:
  finite {i. 0 < (SUP MA. count M i)}
  if A  {} subset_mset.bdd_above A
proof -
  have "{i. Sup ((λX. count X i) ` A) > 0}  (XA. {i. 0 < count X i})"
  proof safe
    fix i assume pos: "(SUP XA. count X i) > 0"
    show "i  (XA. {i. 0 < count X i})"
    proof (rule ccontr)
      assume "i  (XA. {i. 0 < count X i})"
      hence "XA. count X i  0" by (auto simp: count_eq_zero_iff)
      with that have "(SUP XA. count X i)  0"
        by (intro cSup_least bdd_above_multiset_imp_bdd_above_count) auto
      with pos show False by simp
    qed
  qed
  moreover from that have "finite "
    by (rule bdd_above_multiset_imp_finite_support)
  ultimately show "finite {i. Sup ((λX. count X i) ` A) > 0}"
    by (rule finite_subset)
qed

lemma count_Sup_multiset_nonempty:
  count (Sup A) x = (SUP XA. count X x)
  if A  {} subset_mset.bdd_above A
  using that by (simp add: Sup_multiset_def Sup_multiset_in_multiset count_Abs_multiset)

interpretation subset_mset: conditionally_complete_lattice Inf Sup "(∩#)" "(⊆#)" "(⊂#)" "(∪#)"
proof
  fix X :: "'a multiset" and A
  assume "X  A"
  show "Inf A ⊆# X"
    by (metis X  A count_Inf_multiset_nonempty empty_iff image_eqI mset_subset_eqI wellorder_Inf_le1)
next
  fix X :: "'a multiset" and A
  assume nonempty: "A  {}" and le: "Y. Y  A  X ⊆# Y"
  show "X ⊆# Inf A"
  proof (rule mset_subset_eqI)
    fix x
    from nonempty have "count X x  (INF XA. count X x)"
      by (intro cInf_greatest) (auto intro: mset_subset_eq_count le)
    also from nonempty have " = count (Inf A) x" by (simp add: count_Inf_multiset_nonempty)
    finally show "count X x  count (Inf A) x" .
  qed
next
  fix X :: "'a multiset" and A
  assume X: "X  A" and bdd: "subset_mset.bdd_above A"
  show "X ⊆# Sup A"
  proof (rule mset_subset_eqI)
    fix x
    from X have "A  {}" by auto
    have "count X x  (SUP XA. count X x)"
      by (intro cSUP_upper X bdd_above_multiset_imp_bdd_above_count bdd)
    also from count_Sup_multiset_nonempty[OF A  {} bdd]
      have "(SUP XA. count X x) = count (Sup A) x" by simp
    finally show "count X x  count (Sup A) x" .
  qed
next
  fix X :: "'a multiset" and A
  assume nonempty: "A  {}" and ge: "Y. Y  A  Y ⊆# X"
  from ge have bdd: "subset_mset.bdd_above A"
    by blast
  show "Sup A ⊆# X"
  proof (rule mset_subset_eqI)
    fix x
    from count_Sup_multiset_nonempty[OF A  {} bdd]
      have "count (Sup A) x = (SUP XA. count X x)" .
    also from nonempty have "  count X x"
      by (intro cSup_least) (auto intro: mset_subset_eq_count ge)
    finally show "count (Sup A) x  count X x" .
  qed
qed ― ‹FIXME: avoid junk stemming from type class interpretation›

lemma set_mset_Inf:
  assumes "A  {}"
  shows   "set_mset (Inf A) = (XA. set_mset X)"
proof safe
  fix x X assume "x ∈# Inf A" "X  A"
  hence nonempty: "A  {}" by (auto simp: Inf_multiset_empty)
  from x ∈# Inf A have "{#x#} ⊆# Inf A" by auto
  also from X  A have " ⊆# X" by (rule subset_mset.cInf_lower) simp_all
  finally show "x ∈# X" by simp
next
  fix x assume x: "x  (XA. set_mset X)"
  hence "{#x#} ⊆# X" if "X  A" for X using that by auto
  from assms and this have "{#x#} ⊆# Inf A" by (rule subset_mset.cInf_greatest)
  thus "x ∈# Inf A" by simp
qed

lemma in_Inf_multiset_iff:
  assumes "A  {}"
  shows   "x ∈# Inf A  (XA. x ∈# X)"
proof -
  from assms have "set_mset (Inf A) = (XA. set_mset X)" by (rule set_mset_Inf)
  also have "x    (XA. x ∈# X)" by simp
  finally show ?thesis .
qed

lemma in_Inf_multisetD: "x ∈# Inf A  X  A  x ∈# X"
  by (subst (asm) in_Inf_multiset_iff) auto

lemma set_mset_Sup:
  assumes "subset_mset.bdd_above A"
  shows   "set_mset (Sup A) = (XA. set_mset X)"
proof safe
  fix x assume "x ∈# Sup A"
  hence nonempty: "A  {}" by (auto simp: Sup_multiset_empty)
  show "x  (XA. set_mset X)"
  proof (rule ccontr)
    assume x: "x  (XA. set_mset X)"
    have "count X x  count (Sup A) x" if "X  A" for X x
      using that by (intro mset_subset_eq_count subset_mset.cSup_upper assms)
    with x have "X ⊆# Sup A - {#x#}" if "X  A" for X
      using that by (auto simp: subseteq_mset_def algebra_simps not_in_iff)
    hence "Sup A ⊆# Sup A - {#x#}" by (intro subset_mset.cSup_least nonempty)
    with x ∈# Sup A show False
      using mset_subset_diff_self by fastforce
  qed
next
  fix x X assume "x  set_mset X" "X  A"
  hence "{#x#} ⊆# X" by auto
  also have "X ⊆# Sup A" by (intro subset_mset.cSup_upper X  A assms)
  finally show "x  set_mset (Sup A)" by simp
qed

lemma in_Sup_multiset_iff:
  assumes "subset_mset.bdd_above A"
  shows   "x ∈# Sup A  (XA. x ∈# X)"
  by (simp add: assms set_mset_Sup)

lemma in_Sup_multisetD:
  assumes "x ∈# Sup A"
  shows   "XA. x ∈# X"
  using Sup_multiset_unbounded assms in_Sup_multiset_iff by fastforce

interpretation subset_mset: distrib_lattice "(∩#)" "(⊆#)" "(⊂#)" "(∪#)"
proof
  fix A B C :: "'a multiset"
  show "A ∪# (B ∩# C) = A ∪# B ∩# (A ∪# C)"
    by (intro multiset_eqI) simp_all
qed ― ‹FIXME: avoid junk stemming from type class interpretation›


subsubsection ‹Filter (with comprehension syntax)›

text ‹Multiset comprehension›

lift_definition filter_mset :: "('a  bool)  'a multiset  'a multiset"
is "λP M. λx. if P x then M x else 0"
by (rule filter_preserves_multiset)

syntax (ASCII)
  "_MCollect" :: "pttrn  'a multiset  bool  'a multiset"    ((1{#_ :# _./ _#}))
syntax
  "_MCollect" :: "pttrn  'a multiset  bool  'a multiset"    ((1{#_ ∈# _./ _#}))
syntax_consts
  "_MCollect" == filter_mset
translations
  "{#x ∈# M. P#}" == "CONST filter_mset (λx. P) M"

lemma count_filter_mset [simp]:
  "count (filter_mset P M) a = (if P a then count M a else 0)"
  by (simp add: filter_mset.rep_eq)

lemma set_mset_filter [simp]:
  "set_mset (filter_mset P M) = {a  set_mset M. P a}"
  by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_filter_mset) simp

lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}"
  by (rule multiset_eqI) simp

lemma filter_single_mset: "filter_mset P {#x#} = (if P x then {#x#} else {#})"
  by (rule multiset_eqI) simp

lemma filter_union_mset [simp]: "filter_mset P (M + N) = filter_mset P M + filter_mset P N"
  by (rule multiset_eqI) simp

lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N"
  by (rule multiset_eqI) simp

lemma filter_inter_mset [simp]: "filter_mset P (M ∩# N) = filter_mset P M ∩# filter_mset P N"
  by (rule multiset_eqI) simp

lemma filter_sup_mset[simp]: "filter_mset P (A ∪# B) = filter_mset P A ∪# filter_mset P B"
  by (rule multiset_eqI) simp

lemma filter_mset_add_mset [simp]:
   "filter_mset P (add_mset x A) =
     (if P x then add_mset x (filter_mset P A) else filter_mset P A)"
   by (auto simp: multiset_eq_iff)

lemma multiset_filter_subset[simp]: "filter_mset f M ⊆# M"
  by (simp add: mset_subset_eqI)

lemma multiset_filter_mono:
  assumes "A ⊆# B"
  shows "filter_mset f A ⊆# filter_mset f B"
  by (metis assms filter_sup_mset subset_mset.order_iff)

lemma filter_mset_eq_conv:
  "filter_mset P M = N  N ⊆# M  (b∈#N. P b)  (a∈#M - N. ¬ P a)" (is "?P  ?Q")
proof
  assume ?P then show ?Q by auto (simp add: multiset_eq_iff in_diff_count)
next
  assume ?Q
  then obtain Q where M: "M = N + Q"
    by (auto simp add: mset_subset_eq_exists_conv)
  then have MN: "M - N = Q" by simp
  show ?P
  proof (rule multiset_eqI)
    fix a
    from ?Q MN have *: "¬ P a  a ∉# N" "P a  a ∉# Q"
      by auto
    show "count (filter_mset P M) a = count N a"
    proof (cases "a ∈# M")
      case True
      with * show ?thesis
        by (simp add: not_in_iff M)
    next
      case False then have "count M a = 0"
        by (simp add: not_in_iff)
      with M show ?thesis by simp
    qed
  qed
qed

lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x ∈# M. Q x  P x#}"
  by (auto simp: multiset_eq_iff)

lemma
  filter_mset_True[simp]: "{#y ∈# M. True#} = M" and
  filter_mset_False[simp]: "{#y ∈# M. False#} = {#}"
  by (auto simp: multiset_eq_iff)

lemma filter_mset_cong0:
  assumes "x. x ∈# M  f x  g x"
  shows "filter_mset f M = filter_mset g M"
proof (rule subset_mset.antisym; unfold subseteq_mset_def; rule allI)
  fix x
  show "count (filter_mset f M) x  count (filter_mset g M) x"
    using assms by (cases "x ∈# M") (simp_all add: not_in_iff)
next
  fix x
  show "count (filter_mset g M) x  count (filter_mset f M) x"
    using assms by (cases "x ∈# M") (simp_all add: not_in_iff)
qed

lemma filter_mset_cong:
  assumes "M = M'" and "x. x ∈# M'  f x  g x"
  shows "filter_mset f M = filter_mset g M'"
  unfolding M = M'
  using assms by (auto intro: filter_mset_cong0)

lemma filter_eq_replicate_mset: "{#y ∈# D. y = x#} = replicate_mset (count D x) x"
  by (induct D) (simp add: multiset_eqI)


subsubsection ‹Size›

definition wcount where "wcount f M = (λx. count M x * Suc (f x))"

lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a"
  by (auto simp: wcount_def add_mult_distrib)

lemma wcount_add_mset:
  "wcount f (add_mset x M) a = (if x = a then Suc (f a) else 0) + wcount f M a"
  unfolding add_mset_add_single[of _ M] wcount_union by (auto simp: wcount_def)

definition size_multiset :: "('a  nat)  'a multiset  nat" where
  "size_multiset f M = sum (wcount f M) (set_mset M)"

lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def]

instantiation multiset :: (type) size
begin

definition size_multiset where
  size_multiset_overloaded_def: "size_multiset = Multiset.size_multiset (λ_. 0)"
instance ..

end

lemmas size_multiset_overloaded_eq =
  size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified]

lemma size_multiset_empty [simp]: "size_multiset f {#} = 0"
  by (simp add: size_multiset_def)

lemma size_empty [simp]: "size {#} = 0"
  by (simp add: size_multiset_overloaded_def)

lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)"
  by (simp add: size_multiset_eq)

lemma size_single: "size {#b#} = 1"
  by (simp add: size_multiset_overloaded_def size_multiset_single)

lemma sum_wcount_Int:
  "finite A  sum (wcount f N) (A  set_mset N) = sum (wcount f N) A"
  by (induct rule: finite_induct)
    (simp_all add: Int_insert_left wcount_def count_eq_zero_iff)

lemma size_multiset_union [simp]:
  "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N"
  apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union)
  by (metis add_implies_diff finite_set_mset inf.commute sum_wcount_Int)

lemma size_multiset_add_mset [simp]:
  "size_multiset f (add_mset a M) = Suc (f a) + size_multiset f M"
  by (metis add.commute add_mset_add_single size_multiset_single size_multiset_union)

lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)"
  by (simp add: size_multiset_overloaded_def wcount_add_mset)

lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N"
  by (auto simp add: size_multiset_overloaded_def)

lemma size_multiset_eq_0_iff_empty [iff]:
  "size_multiset f M = 0  M = {#}"
  by (auto simp add: size_multiset_eq count_eq_zero_iff)

lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
by (auto simp add: size_multiset_overloaded_def)

lemma nonempty_has_size: "(S  {#}) = (0 < size S)"
  by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty)

lemma size_eq_Suc_imp_elem: "size M = Suc n  a. a ∈# M"
  using all_not_in_conv by fastforce

lemma size_eq_Suc_imp_eq_union:
  assumes "size M = Suc n"
  shows "a N. M = add_mset a N"
  by (metis assms insert_DiffM size_eq_Suc_imp_elem)

lemma size_mset_mono:
  fixes A B :: "'a multiset"
  assumes "A ⊆# B"
  shows "size A  size B"
proof -
  from assms[unfolded mset_subset_eq_exists_conv]
  obtain C where B: "B = A + C" by auto
  show ?thesis unfolding B by (induct C) auto
qed

lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M)  size M"
  by (rule size_mset_mono[OF multiset_filter_subset])

lemma size_Diff_submset:
  "M ⊆# M'  size (M' - M) = size M' - size(M::'a multiset)"
by (metis add_diff_cancel_left' size_union mset_subset_eq_exists_conv)

lemma size_lt_imp_ex_count_lt: "size M < size N  x ∈# N. count M x < count N x"
  by (metis count_eq_zero_iff leD not_le_imp_less not_less_zero size_mset_mono subseteq_mset_def)


subsection ‹Induction and case splits›

theorem multiset_induct [case_names empty add, induct type: multiset]:
  assumes empty: "P {#}"
  assumes add: "x M. P M  P (add_mset x M)"
  shows "P M"
proof (induct "size M" arbitrary: M)
  case 0 thus "P M" by (simp add: empty)
next
  case (Suc k)
  obtain N x where "M = add_mset x N"
    using Suc k = size M [symmetric]
    using size_eq_Suc_imp_eq_union by fast
  with Suc add show "P M" by simp
qed

lemma multiset_induct_min[case_names empty add]:
  fixes M :: "'a::linorder multiset"
  assumes
    empty: "P {#}" and
    add: "x M. P M  (y ∈# M. y  x)  P (add_mset x M)"
  shows "P M"
proof (induct "size M" arbitrary: M)
  case (Suc k)
  note ih = this(1) and Sk_eq_sz_M = this(2)

  let ?y = "Min_mset M"
  let ?N = "M - {#?y#}"

  have M: "M = add_mset ?y ?N"
    by (metis Min_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero
      set_mset_eq_empty_iff size_empty)
  show ?case
    by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset,
      meson Min_le finite_set_mset in_diffD)
qed (simp add: empty)

lemma multiset_induct_max[case_names empty add]:
  fixes M :: "'a::linorder multiset"
  assumes
    empty: "P {#}" and
    add: "x M. P M  (y ∈# M. y  x)  P (add_mset x M)"
  shows "P M"
proof (induct "size M" arbitrary: M)
  case (Suc k)
  note ih = this(1) and Sk_eq_sz_M = this(2)

  let ?y = "Max_mset M"
  let ?N = "M - {#?y#}"

  have M: "M = add_mset ?y ?N"
    by (metis Max_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero
      set_mset_eq_empty_iff size_empty)
  show ?case
    by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset,
      meson Max_ge finite_set_mset in_diffD)
qed (simp add: empty)

lemma multi_nonempty_split: "M  {#}  A a. M = add_mset a A"
  by (induct M) auto

lemma multiset_cases [cases type]:
  obtains (empty) "M = {#}" | (add) x N where "M = add_mset x N"
  by (induct M) simp_all

lemma multi_drop_mem_not_eq: "c ∈# B  B