src/HOL/Library/Multiset.thy
author haftmann
Sat, 14 May 2011 18:26:25 +0200
changeset 43670 5b45125b15ba
parent 41752 6d19301074cf
child 43740 1c0b99f950d9
permissions -rw-r--r--
use pointfree characterisation for fold_set locale
     1 (*  Title:      HOL/Library/Multiset.thy
     2     Author:     Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker
     3 *)
     4 
     5 header {* (Finite) multisets *}
     6 
     7 theory Multiset
     8 imports Main
     9 begin
    10 
    11 subsection {* The type of multisets *}
    12 
    13 typedef 'a multiset = "{f :: 'a => nat. finite {x. f x > 0}}"
    14   morphisms count Abs_multiset
    15 proof
    16   show "(\<lambda>x. 0::nat) \<in> ?multiset" by simp
    17 qed
    18 
    19 lemmas multiset_typedef = Abs_multiset_inverse count_inverse count
    20 
    21 abbreviation Melem :: "'a => 'a multiset => bool"  ("(_/ :# _)" [50, 51] 50) where
    22   "a :# M == 0 < count M a"
    23 
    24 notation (xsymbols)
    25   Melem (infix "\<in>#" 50)
    26 
    27 lemma multiset_eq_iff:
    28   "M = N \<longleftrightarrow> (\<forall>a. count M a = count N a)"
    29   by (simp only: count_inject [symmetric] fun_eq_iff)
    30 
    31 lemma multiset_eqI:
    32   "(\<And>x. count A x = count B x) \<Longrightarrow> A = B"
    33   using multiset_eq_iff by auto
    34 
    35 text {*
    36  \medskip Preservation of the representing set @{term multiset}.
    37 *}
    38 
    39 lemma const0_in_multiset:
    40   "(\<lambda>a. 0) \<in> multiset"
    41   by (simp add: multiset_def)
    42 
    43 lemma only1_in_multiset:
    44   "(\<lambda>b. if b = a then n else 0) \<in> multiset"
    45   by (simp add: multiset_def)
    46 
    47 lemma union_preserves_multiset:
    48   "M \<in> multiset \<Longrightarrow> N \<in> multiset \<Longrightarrow> (\<lambda>a. M a + N a) \<in> multiset"
    49   by (simp add: multiset_def)
    50 
    51 lemma diff_preserves_multiset:
    52   assumes "M \<in> multiset"
    53   shows "(\<lambda>a. M a - N a) \<in> multiset"
    54 proof -
    55   have "{x. N x < M x} \<subseteq> {x. 0 < M x}"
    56     by auto
    57   with assms show ?thesis
    58     by (auto simp add: multiset_def intro: finite_subset)
    59 qed
    60 
    61 lemma filter_preserves_multiset:
    62   assumes "M \<in> multiset"
    63   shows "(\<lambda>x. if P x then M x else 0) \<in> multiset"
    64 proof -
    65   have "{x. (P x \<longrightarrow> 0 < M x) \<and> P x} \<subseteq> {x. 0 < M x}"
    66     by auto
    67   with assms show ?thesis
    68     by (auto simp add: multiset_def intro: finite_subset)
    69 qed
    70 
    71 lemmas in_multiset = const0_in_multiset only1_in_multiset
    72   union_preserves_multiset diff_preserves_multiset filter_preserves_multiset
    73 
    74 
    75 subsection {* Representing multisets *}
    76 
    77 text {* Multiset enumeration *}
    78 
    79 instantiation multiset :: (type) "{zero, plus}"
    80 begin
    81 
    82 definition Mempty_def:
    83   "0 = Abs_multiset (\<lambda>a. 0)"
    84 
    85 abbreviation Mempty :: "'a multiset" ("{#}") where
    86   "Mempty \<equiv> 0"
    87 
    88 definition union_def:
    89   "M + N = Abs_multiset (\<lambda>a. count M a + count N a)"
    90 
    91 instance ..
    92 
    93 end
    94 
    95 definition single :: "'a => 'a multiset" where
    96   "single a = Abs_multiset (\<lambda>b. if b = a then 1 else 0)"
    97 
    98 syntax
    99   "_multiset" :: "args => 'a multiset"    ("{#(_)#}")
   100 translations
   101   "{#x, xs#}" == "{#x#} + {#xs#}"
   102   "{#x#}" == "CONST single x"
   103 
   104 lemma count_empty [simp]: "count {#} a = 0"
   105   by (simp add: Mempty_def in_multiset multiset_typedef)
   106 
   107 lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)"
   108   by (simp add: single_def in_multiset multiset_typedef)
   109 
   110 
   111 subsection {* Basic operations *}
   112 
   113 subsubsection {* Union *}
   114 
   115 lemma count_union [simp]: "count (M + N) a = count M a + count N a"
   116   by (simp add: union_def in_multiset multiset_typedef)
   117 
   118 instance multiset :: (type) cancel_comm_monoid_add proof
   119 qed (simp_all add: multiset_eq_iff)
   120 
   121 
   122 subsubsection {* Difference *}
   123 
   124 instantiation multiset :: (type) minus
   125 begin
   126 
   127 definition diff_def:
   128   "M - N = Abs_multiset (\<lambda>a. count M a - count N a)"
   129 
   130 instance ..
   131 
   132 end
   133 
   134 lemma count_diff [simp]: "count (M - N) a = count M a - count N a"
   135   by (simp add: diff_def in_multiset multiset_typedef)
   136 
   137 lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
   138 by(simp add: multiset_eq_iff)
   139 
   140 lemma diff_cancel[simp]: "A - A = {#}"
   141 by (rule multiset_eqI) simp
   142 
   143 lemma diff_union_cancelR [simp]: "M + N - N = (M::'a multiset)"
   144 by(simp add: multiset_eq_iff)
   145 
   146 lemma diff_union_cancelL [simp]: "N + M - N = (M::'a multiset)"
   147 by(simp add: multiset_eq_iff)
   148 
   149 lemma insert_DiffM:
   150   "x \<in># M \<Longrightarrow> {#x#} + (M - {#x#}) = M"
   151   by (clarsimp simp: multiset_eq_iff)
   152 
   153 lemma insert_DiffM2 [simp]:
   154   "x \<in># M \<Longrightarrow> M - {#x#} + {#x#} = M"
   155   by (clarsimp simp: multiset_eq_iff)
   156 
   157 lemma diff_right_commute:
   158   "(M::'a multiset) - N - Q = M - Q - N"
   159   by (auto simp add: multiset_eq_iff)
   160 
   161 lemma diff_add:
   162   "(M::'a multiset) - (N + Q) = M - N - Q"
   163 by (simp add: multiset_eq_iff)
   164 
   165 lemma diff_union_swap:
   166   "a \<noteq> b \<Longrightarrow> M - {#a#} + {#b#} = M + {#b#} - {#a#}"
   167   by (auto simp add: multiset_eq_iff)
   168 
   169 lemma diff_union_single_conv:
   170   "a \<in># J \<Longrightarrow> I + J - {#a#} = I + (J - {#a#})"
   171   by (simp add: multiset_eq_iff)
   172 
   173 
   174 subsubsection {* Equality of multisets *}
   175 
   176 lemma single_not_empty [simp]: "{#a#} \<noteq> {#} \<and> {#} \<noteq> {#a#}"
   177   by (simp add: multiset_eq_iff)
   178 
   179 lemma single_eq_single [simp]: "{#a#} = {#b#} \<longleftrightarrow> a = b"
   180   by (auto simp add: multiset_eq_iff)
   181 
   182 lemma union_eq_empty [iff]: "M + N = {#} \<longleftrightarrow> M = {#} \<and> N = {#}"
   183   by (auto simp add: multiset_eq_iff)
   184 
   185 lemma empty_eq_union [iff]: "{#} = M + N \<longleftrightarrow> M = {#} \<and> N = {#}"
   186   by (auto simp add: multiset_eq_iff)
   187 
   188 lemma multi_self_add_other_not_self [simp]: "M = M + {#x#} \<longleftrightarrow> False"
   189   by (auto simp add: multiset_eq_iff)
   190 
   191 lemma diff_single_trivial:
   192   "\<not> x \<in># M \<Longrightarrow> M - {#x#} = M"
   193   by (auto simp add: multiset_eq_iff)
   194 
   195 lemma diff_single_eq_union:
   196   "x \<in># M \<Longrightarrow> M - {#x#} = N \<longleftrightarrow> M = N + {#x#}"
   197   by auto
   198 
   199 lemma union_single_eq_diff:
   200   "M + {#x#} = N \<Longrightarrow> M = N - {#x#}"
   201   by (auto dest: sym)
   202 
   203 lemma union_single_eq_member:
   204   "M + {#x#} = N \<Longrightarrow> x \<in># N"
   205   by auto
   206 
   207 lemma union_is_single:
   208   "M + N = {#a#} \<longleftrightarrow> M = {#a#} \<and> N={#} \<or> M = {#} \<and> N = {#a#}" (is "?lhs = ?rhs")proof
   209   assume ?rhs then show ?lhs by auto
   210 next
   211   assume ?lhs thus ?rhs
   212     by(simp add: multiset_eq_iff split:if_splits) (metis add_is_1)
   213 qed
   214 
   215 lemma single_is_union:
   216   "{#a#} = M + N \<longleftrightarrow> {#a#} = M \<and> N = {#} \<or> M = {#} \<and> {#a#} = N"
   217   by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single)
   218 
   219 lemma add_eq_conv_diff:
   220   "M + {#a#} = N + {#b#} \<longleftrightarrow> M = N \<and> a = b \<or> M = N - {#a#} + {#b#} \<and> N = M - {#b#} + {#a#}"  (is "?lhs = ?rhs")
   221 (* shorter: by (simp add: multiset_eq_iff) fastsimp *)
   222 proof
   223   assume ?rhs then show ?lhs
   224   by (auto simp add: add_assoc add_commute [of "{#b#}"])
   225     (drule sym, simp add: add_assoc [symmetric])
   226 next
   227   assume ?lhs
   228   show ?rhs
   229   proof (cases "a = b")
   230     case True with `?lhs` show ?thesis by simp
   231   next
   232     case False
   233     from `?lhs` have "a \<in># N + {#b#}" by (rule union_single_eq_member)
   234     with False have "a \<in># N" by auto
   235     moreover from `?lhs` have "M = N + {#b#} - {#a#}" by (rule union_single_eq_diff)
   236     moreover note False
   237     ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"] diff_union_swap)
   238   qed
   239 qed
   240 
   241 lemma insert_noteq_member: 
   242   assumes BC: "B + {#b#} = C + {#c#}"
   243    and bnotc: "b \<noteq> c"
   244   shows "c \<in># B"
   245 proof -
   246   have "c \<in># C + {#c#}" by simp
   247   have nc: "\<not> c \<in># {#b#}" using bnotc by simp
   248   then have "c \<in># B + {#b#}" using BC by simp
   249   then show "c \<in># B" using nc by simp
   250 qed
   251 
   252 lemma add_eq_conv_ex:
   253   "(M + {#a#} = N + {#b#}) =
   254     (M = N \<and> a = b \<or> (\<exists>K. M = K + {#b#} \<and> N = K + {#a#}))"
   255   by (auto simp add: add_eq_conv_diff)
   256 
   257 
   258 subsubsection {* Pointwise ordering induced by count *}
   259 
   260 instantiation multiset :: (type) ordered_ab_semigroup_add_imp_le
   261 begin
   262 
   263 definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
   264   mset_le_def: "A \<le> B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)"
   265 
   266 definition less_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
   267   mset_less_def: "(A::'a multiset) < B \<longleftrightarrow> A \<le> B \<and> A \<noteq> B"
   268 
   269 instance proof
   270 qed (auto simp add: mset_le_def mset_less_def multiset_eq_iff intro: order_trans antisym)
   271 
   272 end
   273 
   274 lemma mset_less_eqI:
   275   "(\<And>x. count A x \<le> count B x) \<Longrightarrow> A \<le> B"
   276   by (simp add: mset_le_def)
   277 
   278 lemma mset_le_exists_conv:
   279   "(A::'a multiset) \<le> B \<longleftrightarrow> (\<exists>C. B = A + C)"
   280 apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI)
   281 apply (auto intro: multiset_eq_iff [THEN iffD2])
   282 done
   283 
   284 lemma mset_le_mono_add_right_cancel [simp]:
   285   "(A::'a multiset) + C \<le> B + C \<longleftrightarrow> A \<le> B"
   286   by (fact add_le_cancel_right)
   287 
   288 lemma mset_le_mono_add_left_cancel [simp]:
   289   "C + (A::'a multiset) \<le> C + B \<longleftrightarrow> A \<le> B"
   290   by (fact add_le_cancel_left)
   291 
   292 lemma mset_le_mono_add:
   293   "(A::'a multiset) \<le> B \<Longrightarrow> C \<le> D \<Longrightarrow> A + C \<le> B + D"
   294   by (fact add_mono)
   295 
   296 lemma mset_le_add_left [simp]:
   297   "(A::'a multiset) \<le> A + B"
   298   unfolding mset_le_def by auto
   299 
   300 lemma mset_le_add_right [simp]:
   301   "B \<le> (A::'a multiset) + B"
   302   unfolding mset_le_def by auto
   303 
   304 lemma mset_le_single:
   305   "a :# B \<Longrightarrow> {#a#} \<le> B"
   306   by (simp add: mset_le_def)
   307 
   308 lemma multiset_diff_union_assoc:
   309   "C \<le> B \<Longrightarrow> (A::'a multiset) + B - C = A + (B - C)"
   310   by (simp add: multiset_eq_iff mset_le_def)
   311 
   312 lemma mset_le_multiset_union_diff_commute:
   313   "B \<le> A \<Longrightarrow> (A::'a multiset) - B + C = A + C - B"
   314 by (simp add: multiset_eq_iff mset_le_def)
   315 
   316 lemma diff_le_self[simp]: "(M::'a multiset) - N \<le> M"
   317 by(simp add: mset_le_def)
   318 
   319 lemma mset_lessD: "A < B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B"
   320 apply (clarsimp simp: mset_le_def mset_less_def)
   321 apply (erule_tac x=x in allE)
   322 apply auto
   323 done
   324 
   325 lemma mset_leD: "A \<le> B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B"
   326 apply (clarsimp simp: mset_le_def mset_less_def)
   327 apply (erule_tac x = x in allE)
   328 apply auto
   329 done
   330   
   331 lemma mset_less_insertD: "(A + {#x#} < B) \<Longrightarrow> (x \<in># B \<and> A < B)"
   332 apply (rule conjI)
   333  apply (simp add: mset_lessD)
   334 apply (clarsimp simp: mset_le_def mset_less_def)
   335 apply safe
   336  apply (erule_tac x = a in allE)
   337  apply (auto split: split_if_asm)
   338 done
   339 
   340 lemma mset_le_insertD: "(A + {#x#} \<le> B) \<Longrightarrow> (x \<in># B \<and> A \<le> B)"
   341 apply (rule conjI)
   342  apply (simp add: mset_leD)
   343 apply (force simp: mset_le_def mset_less_def split: split_if_asm)
   344 done
   345 
   346 lemma mset_less_of_empty[simp]: "A < {#} \<longleftrightarrow> False"
   347   by (auto simp add: mset_less_def mset_le_def multiset_eq_iff)
   348 
   349 lemma multi_psub_of_add_self[simp]: "A < A + {#x#}"
   350   by (auto simp: mset_le_def mset_less_def)
   351 
   352 lemma multi_psub_self[simp]: "(A::'a multiset) < A = False"
   353   by simp
   354 
   355 lemma mset_less_add_bothsides:
   356   "T + {#x#} < S + {#x#} \<Longrightarrow> T < S"
   357   by (fact add_less_imp_less_right)
   358 
   359 lemma mset_less_empty_nonempty:
   360   "{#} < S \<longleftrightarrow> S \<noteq> {#}"
   361   by (auto simp: mset_le_def mset_less_def)
   362 
   363 lemma mset_less_diff_self:
   364   "c \<in># B \<Longrightarrow> B - {#c#} < B"
   365   by (auto simp: mset_le_def mset_less_def multiset_eq_iff)
   366 
   367 
   368 subsubsection {* Intersection *}
   369 
   370 instantiation multiset :: (type) semilattice_inf
   371 begin
   372 
   373 definition inf_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where
   374   multiset_inter_def: "inf_multiset A B = A - (A - B)"
   375 
   376 instance proof -
   377   have aux: "\<And>m n q :: nat. m \<le> n \<Longrightarrow> m \<le> q \<Longrightarrow> m \<le> n - (n - q)" by arith
   378   show "OFCLASS('a multiset, semilattice_inf_class)" proof
   379   qed (auto simp add: multiset_inter_def mset_le_def aux)
   380 qed
   381 
   382 end
   383 
   384 abbreviation multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where
   385   "multiset_inter \<equiv> inf"
   386 
   387 lemma multiset_inter_count [simp]:
   388   "count (A #\<inter> B) x = min (count A x) (count B x)"
   389   by (simp add: multiset_inter_def multiset_typedef)
   390 
   391 lemma multiset_inter_single: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
   392   by (rule multiset_eqI) (auto simp add: multiset_inter_count)
   393 
   394 lemma multiset_union_diff_commute:
   395   assumes "B #\<inter> C = {#}"
   396   shows "A + B - C = A - C + B"
   397 proof (rule multiset_eqI)
   398   fix x
   399   from assms have "min (count B x) (count C x) = 0"
   400     by (auto simp add: multiset_inter_count multiset_eq_iff)
   401   then have "count B x = 0 \<or> count C x = 0"
   402     by auto
   403   then show "count (A + B - C) x = count (A - C + B) x"
   404     by auto
   405 qed
   406 
   407 
   408 subsubsection {* Filter (with comprehension syntax) *}
   409 
   410 text {* Multiset comprehension *}
   411 
   412 definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where
   413   "filter P M = Abs_multiset (\<lambda>x. if P x then count M x else 0)"
   414 
   415 hide_const (open) filter
   416 
   417 lemma count_filter [simp]:
   418   "count (Multiset.filter P M) a = (if P a then count M a else 0)"
   419   by (simp add: filter_def in_multiset multiset_typedef)
   420 
   421 lemma filter_empty [simp]:
   422   "Multiset.filter P {#} = {#}"
   423   by (rule multiset_eqI) simp
   424 
   425 lemma filter_single [simp]:
   426   "Multiset.filter P {#x#} = (if P x then {#x#} else {#})"
   427   by (rule multiset_eqI) simp
   428 
   429 lemma filter_union [simp]:
   430   "Multiset.filter P (M + N) = Multiset.filter P M + Multiset.filter P N"
   431   by (rule multiset_eqI) simp
   432 
   433 lemma filter_diff [simp]:
   434   "Multiset.filter P (M - N) = Multiset.filter P M - Multiset.filter P N"
   435   by (rule multiset_eqI) simp
   436 
   437 lemma filter_inter [simp]:
   438   "Multiset.filter P (M #\<inter> N) = Multiset.filter P M #\<inter> Multiset.filter P N"
   439   by (rule multiset_eqI) simp
   440 
   441 syntax
   442   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{# _ :# _./ _#})")
   443 syntax (xsymbol)
   444   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{# _ \<in># _./ _#})")
   445 translations
   446   "{#x \<in># M. P#}" == "CONST Multiset.filter (\<lambda>x. P) M"
   447 
   448 
   449 subsubsection {* Set of elements *}
   450 
   451 definition set_of :: "'a multiset => 'a set" where
   452   "set_of M = {x. x :# M}"
   453 
   454 lemma set_of_empty [simp]: "set_of {#} = {}"
   455 by (simp add: set_of_def)
   456 
   457 lemma set_of_single [simp]: "set_of {#b#} = {b}"
   458 by (simp add: set_of_def)
   459 
   460 lemma set_of_union [simp]: "set_of (M + N) = set_of M \<union> set_of N"
   461 by (auto simp add: set_of_def)
   462 
   463 lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})"
   464 by (auto simp add: set_of_def multiset_eq_iff)
   465 
   466 lemma mem_set_of_iff [simp]: "(x \<in> set_of M) = (x :# M)"
   467 by (auto simp add: set_of_def)
   468 
   469 lemma set_of_filter [simp]: "set_of {# x:#M. P x #} = set_of M \<inter> {x. P x}"
   470 by (auto simp add: set_of_def)
   471 
   472 lemma finite_set_of [iff]: "finite (set_of M)"
   473   using count [of M] by (simp add: multiset_def set_of_def)
   474 
   475 
   476 subsubsection {* Size *}
   477 
   478 instantiation multiset :: (type) size
   479 begin
   480 
   481 definition size_def:
   482   "size M = setsum (count M) (set_of M)"
   483 
   484 instance ..
   485 
   486 end
   487 
   488 lemma size_empty [simp]: "size {#} = 0"
   489 by (simp add: size_def)
   490 
   491 lemma size_single [simp]: "size {#b#} = 1"
   492 by (simp add: size_def)
   493 
   494 lemma setsum_count_Int:
   495   "finite A ==> setsum (count N) (A \<inter> set_of N) = setsum (count N) A"
   496 apply (induct rule: finite_induct)
   497  apply simp
   498 apply (simp add: Int_insert_left set_of_def)
   499 done
   500 
   501 lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N"
   502 apply (unfold size_def)
   503 apply (subgoal_tac "count (M + N) = (\<lambda>a. count M a + count N a)")
   504  prefer 2
   505  apply (rule ext, simp)
   506 apply (simp (no_asm_simp) add: setsum_Un_nat setsum_addf setsum_count_Int)
   507 apply (subst Int_commute)
   508 apply (simp (no_asm_simp) add: setsum_count_Int)
   509 done
   510 
   511 lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
   512 by (auto simp add: size_def multiset_eq_iff)
   513 
   514 lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
   515 by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty)
   516 
   517 lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \<exists>a. a :# M"
   518 apply (unfold size_def)
   519 apply (drule setsum_SucD)
   520 apply auto
   521 done
   522 
   523 lemma size_eq_Suc_imp_eq_union:
   524   assumes "size M = Suc n"
   525   shows "\<exists>a N. M = N + {#a#}"
   526 proof -
   527   from assms obtain a where "a \<in># M"
   528     by (erule size_eq_Suc_imp_elem [THEN exE])
   529   then have "M = M - {#a#} + {#a#}" by simp
   530   then show ?thesis by blast
   531 qed
   532 
   533 
   534 subsection {* Induction and case splits *}
   535 
   536 lemma setsum_decr:
   537   "finite F ==> (0::nat) < f a ==>
   538     setsum (f (a := f a - 1)) F = (if a\<in>F then setsum f F - 1 else setsum f F)"
   539 apply (induct rule: finite_induct)
   540  apply auto
   541 apply (drule_tac a = a in mk_disjoint_insert, auto)
   542 done
   543 
   544 lemma rep_multiset_induct_aux:
   545 assumes 1: "P (\<lambda>a. (0::nat))"
   546   and 2: "!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))"
   547 shows "\<forall>f. f \<in> multiset --> setsum f {x. f x \<noteq> 0} = n --> P f"
   548 apply (unfold multiset_def)
   549 apply (induct_tac n, simp, clarify)
   550  apply (subgoal_tac "f = (\<lambda>a.0)")
   551   apply simp
   552   apply (rule 1)
   553  apply (rule ext, force, clarify)
   554 apply (frule setsum_SucD, clarify)
   555 apply (rename_tac a)
   556 apply (subgoal_tac "finite {x. (f (a := f a - 1)) x > 0}")
   557  prefer 2
   558  apply (rule finite_subset)
   559   prefer 2
   560   apply assumption
   561  apply simp
   562  apply blast
   563 apply (subgoal_tac "f = (f (a := f a - 1))(a := (f (a := f a - 1)) a + 1)")
   564  prefer 2
   565  apply (rule ext)
   566  apply (simp (no_asm_simp))
   567  apply (erule ssubst, rule 2 [unfolded multiset_def], blast)
   568 apply (erule allE, erule impE, erule_tac [2] mp, blast)
   569 apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply One_nat_def)
   570 apply (subgoal_tac "{x. x \<noteq> a --> f x \<noteq> 0} = {x. f x \<noteq> 0}")
   571  prefer 2
   572  apply blast
   573 apply (subgoal_tac "{x. x \<noteq> a \<and> f x \<noteq> 0} = {x. f x \<noteq> 0} - {a}")
   574  prefer 2
   575  apply blast
   576 apply (simp add: le_imp_diff_is_add setsum_diff1_nat cong: conj_cong)
   577 done
   578 
   579 theorem rep_multiset_induct:
   580   "f \<in> multiset ==> P (\<lambda>a. 0) ==>
   581     (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))) ==> P f"
   582 using rep_multiset_induct_aux by blast
   583 
   584 theorem multiset_induct [case_names empty add, induct type: multiset]:
   585 assumes empty: "P {#}"
   586   and add: "!!M x. P M ==> P (M + {#x#})"
   587 shows "P M"
   588 proof -
   589   note defns = union_def single_def Mempty_def
   590   note add' = add [unfolded defns, simplified]
   591   have aux: "\<And>a::'a. count (Abs_multiset (\<lambda>b. if b = a then 1 else 0)) =
   592     (\<lambda>b. if b = a then 1 else 0)" by (simp add: Abs_multiset_inverse in_multiset) 
   593   show ?thesis
   594     apply (rule count_inverse [THEN subst])
   595     apply (rule count [THEN rep_multiset_induct])
   596      apply (rule empty [unfolded defns])
   597     apply (subgoal_tac "f(b := f b + 1) = (\<lambda>a. f a + (if a=b then 1 else 0))")
   598      prefer 2
   599      apply (simp add: fun_eq_iff)
   600     apply (erule ssubst)
   601     apply (erule Abs_multiset_inverse [THEN subst])
   602     apply (drule add')
   603     apply (simp add: aux)
   604     done
   605 qed
   606 
   607 lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = A + {#a#}"
   608 by (induct M) auto
   609 
   610 lemma multiset_cases [cases type, case_names empty add]:
   611 assumes em:  "M = {#} \<Longrightarrow> P"
   612 assumes add: "\<And>N x. M = N + {#x#} \<Longrightarrow> P"
   613 shows "P"
   614 proof (cases "M = {#}")
   615   assume "M = {#}" then show ?thesis using em by simp
   616 next
   617   assume "M \<noteq> {#}"
   618   then obtain M' m where "M = M' + {#m#}" 
   619     by (blast dest: multi_nonempty_split)
   620   then show ?thesis using add by simp
   621 qed
   622 
   623 lemma multi_member_split: "x \<in># M \<Longrightarrow> \<exists>A. M = A + {#x#}"
   624 apply (cases M)
   625  apply simp
   626 apply (rule_tac x="M - {#x#}" in exI, simp)
   627 done
   628 
   629 lemma multi_drop_mem_not_eq: "c \<in># B \<Longrightarrow> B - {#c#} \<noteq> B"
   630 by (cases "B = {#}") (auto dest: multi_member_split)
   631 
   632 lemma multiset_partition: "M = {# x:#M. P x #} + {# x:#M. \<not> P x #}"
   633 apply (subst multiset_eq_iff)
   634 apply auto
   635 done
   636 
   637 lemma mset_less_size: "(A::'a multiset) < B \<Longrightarrow> size A < size B"
   638 proof (induct A arbitrary: B)
   639   case (empty M)
   640   then have "M \<noteq> {#}" by (simp add: mset_less_empty_nonempty)
   641   then obtain M' x where "M = M' + {#x#}" 
   642     by (blast dest: multi_nonempty_split)
   643   then show ?case by simp
   644 next
   645   case (add S x T)
   646   have IH: "\<And>B. S < B \<Longrightarrow> size S < size B" by fact
   647   have SxsubT: "S + {#x#} < T" by fact
   648   then have "x \<in># T" and "S < T" by (auto dest: mset_less_insertD)
   649   then obtain T' where T: "T = T' + {#x#}" 
   650     by (blast dest: multi_member_split)
   651   then have "S < T'" using SxsubT 
   652     by (blast intro: mset_less_add_bothsides)
   653   then have "size S < size T'" using IH by simp
   654   then show ?case using T by simp
   655 qed
   656 
   657 
   658 subsubsection {* Strong induction and subset induction for multisets *}
   659 
   660 text {* Well-foundedness of proper subset operator: *}
   661 
   662 text {* proper multiset subset *}
   663 
   664 definition
   665   mset_less_rel :: "('a multiset * 'a multiset) set" where
   666   "mset_less_rel = {(A,B). A < B}"
   667 
   668 lemma multiset_add_sub_el_shuffle: 
   669   assumes "c \<in># B" and "b \<noteq> c" 
   670   shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}"
   671 proof -
   672   from `c \<in># B` obtain A where B: "B = A + {#c#}" 
   673     by (blast dest: multi_member_split)
   674   have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp
   675   then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" 
   676     by (simp add: add_ac)
   677   then show ?thesis using B by simp
   678 qed
   679 
   680 lemma wf_mset_less_rel: "wf mset_less_rel"
   681 apply (unfold mset_less_rel_def)
   682 apply (rule wf_measure [THEN wf_subset, where f1=size])
   683 apply (clarsimp simp: measure_def inv_image_def mset_less_size)
   684 done
   685 
   686 text {* The induction rules: *}
   687 
   688 lemma full_multiset_induct [case_names less]:
   689 assumes ih: "\<And>B. \<forall>(A::'a multiset). A < B \<longrightarrow> P A \<Longrightarrow> P B"
   690 shows "P B"
   691 apply (rule wf_mset_less_rel [THEN wf_induct])
   692 apply (rule ih, auto simp: mset_less_rel_def)
   693 done
   694 
   695 lemma multi_subset_induct [consumes 2, case_names empty add]:
   696 assumes "F \<le> A"
   697   and empty: "P {#}"
   698   and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (F + {#a#})"
   699 shows "P F"
   700 proof -
   701   from `F \<le> A`
   702   show ?thesis
   703   proof (induct F)
   704     show "P {#}" by fact
   705   next
   706     fix x F
   707     assume P: "F \<le> A \<Longrightarrow> P F" and i: "F + {#x#} \<le> A"
   708     show "P (F + {#x#})"
   709     proof (rule insert)
   710       from i show "x \<in># A" by (auto dest: mset_le_insertD)
   711       from i have "F \<le> A" by (auto dest: mset_le_insertD)
   712       with P show "P F" .
   713     qed
   714   qed
   715 qed
   716 
   717 
   718 subsection {* Alternative representations *}
   719 
   720 subsubsection {* Lists *}
   721 
   722 primrec multiset_of :: "'a list \<Rightarrow> 'a multiset" where
   723   "multiset_of [] = {#}" |
   724   "multiset_of (a # x) = multiset_of x + {# a #}"
   725 
   726 lemma in_multiset_in_set:
   727   "x \<in># multiset_of xs \<longleftrightarrow> x \<in> set xs"
   728   by (induct xs) simp_all
   729 
   730 lemma count_multiset_of:
   731   "count (multiset_of xs) x = length (filter (\<lambda>y. x = y) xs)"
   732   by (induct xs) simp_all
   733 
   734 lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])"
   735 by (induct x) auto
   736 
   737 lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])"
   738 by (induct x) auto
   739 
   740 lemma set_of_multiset_of[simp]: "set_of (multiset_of x) = set x"
   741 by (induct x) auto
   742 
   743 lemma mem_set_multiset_eq: "x \<in> set xs = (x :# multiset_of xs)"
   744 by (induct xs) auto
   745 
   746 lemma multiset_of_append [simp]:
   747   "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys"
   748   by (induct xs arbitrary: ys) (auto simp: add_ac)
   749 
   750 lemma multiset_of_filter:
   751   "multiset_of (filter P xs) = {#x :# multiset_of xs. P x #}"
   752   by (induct xs) simp_all
   753 
   754 lemma multiset_of_rev [simp]:
   755   "multiset_of (rev xs) = multiset_of xs"
   756   by (induct xs) simp_all
   757 
   758 lemma surj_multiset_of: "surj multiset_of"
   759 apply (unfold surj_def)
   760 apply (rule allI)
   761 apply (rule_tac M = y in multiset_induct)
   762  apply auto
   763 apply (rule_tac x = "x # xa" in exI)
   764 apply auto
   765 done
   766 
   767 lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}"
   768 by (induct x) auto
   769 
   770 lemma distinct_count_atmost_1:
   771   "distinct x = (! a. count (multiset_of x) a = (if a \<in> set x then 1 else 0))"
   772 apply (induct x, simp, rule iffI, simp_all)
   773 apply (rule conjI)
   774 apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of)
   775 apply (erule_tac x = a in allE, simp, clarify)
   776 apply (erule_tac x = aa in allE, simp)
   777 done
   778 
   779 lemma multiset_of_eq_setD:
   780   "multiset_of xs = multiset_of ys \<Longrightarrow> set xs = set ys"
   781 by (rule) (auto simp add:multiset_eq_iff set_count_greater_0)
   782 
   783 lemma set_eq_iff_multiset_of_eq_distinct:
   784   "distinct x \<Longrightarrow> distinct y \<Longrightarrow>
   785     (set x = set y) = (multiset_of x = multiset_of y)"
   786 by (auto simp: multiset_eq_iff distinct_count_atmost_1)
   787 
   788 lemma set_eq_iff_multiset_of_remdups_eq:
   789    "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))"
   790 apply (rule iffI)
   791 apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1])
   792 apply (drule distinct_remdups [THEN distinct_remdups
   793       [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]])
   794 apply simp
   795 done
   796 
   797 lemma multiset_of_compl_union [simp]:
   798   "multiset_of [x\<leftarrow>xs. P x] + multiset_of [x\<leftarrow>xs. \<not>P x] = multiset_of xs"
   799   by (induct xs) (auto simp: add_ac)
   800 
   801 lemma count_multiset_of_length_filter:
   802   "count (multiset_of xs) x = length (filter (\<lambda>y. x = y) xs)"
   803   by (induct xs) auto
   804 
   805 lemma nth_mem_multiset_of: "i < length ls \<Longrightarrow> (ls ! i) :# multiset_of ls"
   806 apply (induct ls arbitrary: i)
   807  apply simp
   808 apply (case_tac i)
   809  apply auto
   810 done
   811 
   812 lemma multiset_of_remove1[simp]:
   813   "multiset_of (remove1 a xs) = multiset_of xs - {#a#}"
   814 by (induct xs) (auto simp add: multiset_eq_iff)
   815 
   816 lemma multiset_of_eq_length:
   817   assumes "multiset_of xs = multiset_of ys"
   818   shows "length xs = length ys"
   819 using assms proof (induct xs arbitrary: ys)
   820   case Nil then show ?case by simp
   821 next
   822   case (Cons x xs)
   823   then have "x \<in># multiset_of ys" by (simp add: union_single_eq_member)
   824   then have "x \<in> set ys" by (simp add: in_multiset_in_set)
   825   from Cons.prems [symmetric] have "multiset_of xs = multiset_of (remove1 x ys)"
   826     by simp
   827   with Cons.hyps have "length xs = length (remove1 x ys)" .
   828   with `x \<in> set ys` show ?case
   829     by (auto simp add: length_remove1 dest: length_pos_if_in_set)
   830 qed
   831 
   832 lemma multiset_of_eq_length_filter:
   833   assumes "multiset_of xs = multiset_of ys"
   834   shows "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) ys)"
   835 proof (cases "z \<in># multiset_of xs")
   836   case False
   837   moreover have "\<not> z \<in># multiset_of ys" using assms False by simp
   838   ultimately show ?thesis by (simp add: count_multiset_of_length_filter)
   839 next
   840   case True
   841   moreover have "z \<in># multiset_of ys" using assms True by simp
   842   show ?thesis using assms proof (induct xs arbitrary: ys)
   843     case Nil then show ?case by simp
   844   next
   845     case (Cons x xs)
   846     from `multiset_of (x # xs) = multiset_of ys` [symmetric]
   847       have *: "multiset_of xs = multiset_of (remove1 x ys)"
   848       and "x \<in> set ys"
   849       by (auto simp add: mem_set_multiset_eq)
   850     from * have "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) (remove1 x ys))" by (rule Cons.hyps)
   851     moreover from `x \<in> set ys` have "length (filter (\<lambda>y. x = y) ys) > 0" by (simp add: filter_empty_conv)
   852     ultimately show ?case using `x \<in> set ys`
   853       by (simp add: filter_remove1) (auto simp add: length_remove1)
   854   qed
   855 qed
   856 
   857 context linorder
   858 begin
   859 
   860 lemma multiset_of_insort [simp]:
   861   "multiset_of (insort_key k x xs) = {#x#} + multiset_of xs"
   862   by (induct xs) (simp_all add: ac_simps)
   863  
   864 lemma multiset_of_sort [simp]:
   865   "multiset_of (sort_key k xs) = multiset_of xs"
   866   by (induct xs) (simp_all add: ac_simps)
   867 
   868 text {*
   869   This lemma shows which properties suffice to show that a function
   870   @{text "f"} with @{text "f xs = ys"} behaves like sort.
   871 *}
   872 
   873 lemma properties_for_sort_key:
   874   assumes "multiset_of ys = multiset_of xs"
   875   and "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>x. f k = f x) ys = filter (\<lambda>x. f k = f x) xs"
   876   and "sorted (map f ys)"
   877   shows "sort_key f xs = ys"
   878 using assms proof (induct xs arbitrary: ys)
   879   case Nil then show ?case by simp
   880 next
   881   case (Cons x xs)
   882   from Cons.prems(2) have
   883     "\<forall>k \<in> set ys. filter (\<lambda>x. f k = f x) (remove1 x ys) = filter (\<lambda>x. f k = f x) xs"
   884     by (simp add: filter_remove1)
   885   with Cons.prems have "sort_key f xs = remove1 x ys"
   886     by (auto intro!: Cons.hyps simp add: sorted_map_remove1)
   887   moreover from Cons.prems have "x \<in> set ys"
   888     by (auto simp add: mem_set_multiset_eq intro!: ccontr)
   889   ultimately show ?case using Cons.prems by (simp add: insort_key_remove1)
   890 qed
   891 
   892 lemma properties_for_sort:
   893   assumes multiset: "multiset_of ys = multiset_of xs"
   894   and "sorted ys"
   895   shows "sort xs = ys"
   896 proof (rule properties_for_sort_key)
   897   from multiset show "multiset_of ys = multiset_of xs" .
   898   from `sorted ys` show "sorted (map (\<lambda>x. x) ys)" by simp
   899   from multiset have "\<And>k. length (filter (\<lambda>y. k = y) ys) = length (filter (\<lambda>x. k = x) xs)"
   900     by (rule multiset_of_eq_length_filter)
   901   then have "\<And>k. replicate (length (filter (\<lambda>y. k = y) ys)) k = replicate (length (filter (\<lambda>x. k = x) xs)) k"
   902     by simp
   903   then show "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>y. k = y) ys = filter (\<lambda>x. k = x) xs"
   904     by (simp add: replicate_length_filter)
   905 qed
   906 
   907 lemma sort_key_by_quicksort:
   908   "sort_key f xs = sort_key f [x\<leftarrow>xs. f x < f (xs ! (length xs div 2))]
   909     @ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))]
   910     @ sort_key f [x\<leftarrow>xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs")
   911 proof (rule properties_for_sort_key)
   912   show "multiset_of ?rhs = multiset_of ?lhs"
   913     by (rule multiset_eqI) (auto simp add: multiset_of_filter)
   914 next
   915   show "sorted (map f ?rhs)"
   916     by (auto simp add: sorted_append intro: sorted_map_same)
   917 next
   918   fix l
   919   assume "l \<in> set ?rhs"
   920   let ?pivot = "f (xs ! (length xs div 2))"
   921   have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto
   922   have "[x \<leftarrow> sort_key f xs . f x = f l] = [x \<leftarrow> xs. f x = f l]"
   923     unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same)
   924   with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp
   925   have "\<And>x P. P (f x) ?pivot \<and> f l = f x \<longleftrightarrow> P (f l) ?pivot \<and> f l = f x" by auto
   926   then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] =
   927     [x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp
   928   note *** = this [of "op <"] this [of "op >"] this [of "op ="]
   929   show "[x \<leftarrow> ?rhs. f l = f x] = [x \<leftarrow> ?lhs. f l = f x]"
   930   proof (cases "f l" ?pivot rule: linorder_cases)
   931     case less then moreover have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto
   932     ultimately show ?thesis
   933       by (simp add: filter_sort [symmetric] ** ***)
   934   next
   935     case equal then show ?thesis
   936       by (simp add: * less_le)
   937   next
   938     case greater then moreover have "f l \<noteq> ?pivot" and "\<not> f l < ?pivot" by auto
   939     ultimately show ?thesis
   940       by (simp add: filter_sort [symmetric] ** ***)
   941   qed
   942 qed
   943 
   944 lemma sort_by_quicksort:
   945   "sort xs = sort [x\<leftarrow>xs. x < xs ! (length xs div 2)]
   946     @ [x\<leftarrow>xs. x = xs ! (length xs div 2)]
   947     @ sort [x\<leftarrow>xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs")
   948   using sort_key_by_quicksort [of "\<lambda>x. x", symmetric] by simp
   949 
   950 text {* A stable parametrized quicksort *}
   951 
   952 definition part :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'b list \<times> 'b list \<times> 'b list" where
   953   "part f pivot xs = ([x \<leftarrow> xs. f x < pivot], [x \<leftarrow> xs. f x = pivot], [x \<leftarrow> xs. pivot < f x])"
   954 
   955 lemma part_code [code]:
   956   "part f pivot [] = ([], [], [])"
   957   "part f pivot (x # xs) = (let (lts, eqs, gts) = part f pivot xs; x' = f x in
   958      if x' < pivot then (x # lts, eqs, gts)
   959      else if x' > pivot then (lts, eqs, x # gts)
   960      else (lts, x # eqs, gts))"
   961   by (auto simp add: part_def Let_def split_def)
   962 
   963 lemma sort_key_by_quicksort_code [code]:
   964   "sort_key f xs = (case xs of [] \<Rightarrow> []
   965     | [x] \<Rightarrow> xs
   966     | [x, y] \<Rightarrow> (if f x \<le> f y then xs else [y, x])
   967     | _ \<Rightarrow> (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs
   968        in sort_key f lts @ eqs @ sort_key f gts))"
   969 proof (cases xs)
   970   case Nil then show ?thesis by simp
   971 next
   972   case (Cons _ ys) note hyps = Cons show ?thesis proof (cases ys)
   973     case Nil with hyps show ?thesis by simp
   974   next
   975     case (Cons _ zs) note hyps = hyps Cons show ?thesis proof (cases zs)
   976       case Nil with hyps show ?thesis by auto
   977     next
   978       case Cons 
   979       from sort_key_by_quicksort [of f xs]
   980       have "sort_key f xs = (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs
   981         in sort_key f lts @ eqs @ sort_key f gts)"
   982       by (simp only: split_def Let_def part_def fst_conv snd_conv)
   983       with hyps Cons show ?thesis by (simp only: list.cases)
   984     qed
   985   qed
   986 qed
   987 
   988 end
   989 
   990 hide_const (open) part
   991 
   992 lemma multiset_of_remdups_le: "multiset_of (remdups xs) \<le> multiset_of xs"
   993   by (induct xs) (auto intro: order_trans)
   994 
   995 lemma multiset_of_update:
   996   "i < length ls \<Longrightarrow> multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}"
   997 proof (induct ls arbitrary: i)
   998   case Nil then show ?case by simp
   999 next
  1000   case (Cons x xs)
  1001   show ?case
  1002   proof (cases i)
  1003     case 0 then show ?thesis by simp
  1004   next
  1005     case (Suc i')
  1006     with Cons show ?thesis
  1007       apply simp
  1008       apply (subst add_assoc)
  1009       apply (subst add_commute [of "{#v#}" "{#x#}"])
  1010       apply (subst add_assoc [symmetric])
  1011       apply simp
  1012       apply (rule mset_le_multiset_union_diff_commute)
  1013       apply (simp add: mset_le_single nth_mem_multiset_of)
  1014       done
  1015   qed
  1016 qed
  1017 
  1018 lemma multiset_of_swap:
  1019   "i < length ls \<Longrightarrow> j < length ls \<Longrightarrow>
  1020     multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls"
  1021   by (cases "i = j") (simp_all add: multiset_of_update nth_mem_multiset_of)
  1022 
  1023 
  1024 subsubsection {* Association lists -- including rudimentary code generation *}
  1025 
  1026 definition count_of :: "('a \<times> nat) list \<Rightarrow> 'a \<Rightarrow> nat" where
  1027   "count_of xs x = (case map_of xs x of None \<Rightarrow> 0 | Some n \<Rightarrow> n)"
  1028 
  1029 lemma count_of_multiset:
  1030   "count_of xs \<in> multiset"
  1031 proof -
  1032   let ?A = "{x::'a. 0 < (case map_of xs x of None \<Rightarrow> 0\<Colon>nat | Some (n\<Colon>nat) \<Rightarrow> n)}"
  1033   have "?A \<subseteq> dom (map_of xs)"
  1034   proof
  1035     fix x
  1036     assume "x \<in> ?A"
  1037     then have "0 < (case map_of xs x of None \<Rightarrow> 0\<Colon>nat | Some (n\<Colon>nat) \<Rightarrow> n)" by simp
  1038     then have "map_of xs x \<noteq> None" by (cases "map_of xs x") auto
  1039     then show "x \<in> dom (map_of xs)" by auto
  1040   qed
  1041   with finite_dom_map_of [of xs] have "finite ?A"
  1042     by (auto intro: finite_subset)
  1043   then show ?thesis
  1044     by (simp add: count_of_def fun_eq_iff multiset_def)
  1045 qed
  1046 
  1047 lemma count_simps [simp]:
  1048   "count_of [] = (\<lambda>_. 0)"
  1049   "count_of ((x, n) # xs) = (\<lambda>y. if x = y then n else count_of xs y)"
  1050   by (simp_all add: count_of_def fun_eq_iff)
  1051 
  1052 lemma count_of_empty:
  1053   "x \<notin> fst ` set xs \<Longrightarrow> count_of xs x = 0"
  1054   by (induct xs) (simp_all add: count_of_def)
  1055 
  1056 lemma count_of_filter:
  1057   "count_of (filter (P \<circ> fst) xs) x = (if P x then count_of xs x else 0)"
  1058   by (induct xs) auto
  1059 
  1060 definition Bag :: "('a \<times> nat) list \<Rightarrow> 'a multiset" where
  1061   "Bag xs = Abs_multiset (count_of xs)"
  1062 
  1063 code_datatype Bag
  1064 
  1065 lemma count_Bag [simp, code]:
  1066   "count (Bag xs) = count_of xs"
  1067   by (simp add: Bag_def count_of_multiset Abs_multiset_inverse)
  1068 
  1069 lemma Mempty_Bag [code]:
  1070   "{#} = Bag []"
  1071   by (simp add: multiset_eq_iff)
  1072   
  1073 lemma single_Bag [code]:
  1074   "{#x#} = Bag [(x, 1)]"
  1075   by (simp add: multiset_eq_iff)
  1076 
  1077 lemma filter_Bag [code]:
  1078   "Multiset.filter P (Bag xs) = Bag (filter (P \<circ> fst) xs)"
  1079   by (rule multiset_eqI) (simp add: count_of_filter)
  1080 
  1081 lemma mset_less_eq_Bag [code]:
  1082   "Bag xs \<le> A \<longleftrightarrow> (\<forall>(x, n) \<in> set xs. count_of xs x \<le> count A x)"
  1083     (is "?lhs \<longleftrightarrow> ?rhs")
  1084 proof
  1085   assume ?lhs then show ?rhs
  1086     by (auto simp add: mset_le_def count_Bag)
  1087 next
  1088   assume ?rhs
  1089   show ?lhs
  1090   proof (rule mset_less_eqI)
  1091     fix x
  1092     from `?rhs` have "count_of xs x \<le> count A x"
  1093       by (cases "x \<in> fst ` set xs") (auto simp add: count_of_empty)
  1094     then show "count (Bag xs) x \<le> count A x"
  1095       by (simp add: mset_le_def count_Bag)
  1096   qed
  1097 qed
  1098 
  1099 instantiation multiset :: (equal) equal
  1100 begin
  1101 
  1102 definition
  1103   "HOL.equal A B \<longleftrightarrow> (A::'a multiset) \<le> B \<and> B \<le> A"
  1104 
  1105 instance proof
  1106 qed (simp add: equal_multiset_def eq_iff)
  1107 
  1108 end
  1109 
  1110 lemma [code nbe]:
  1111   "HOL.equal (A :: 'a::equal multiset) A \<longleftrightarrow> True"
  1112   by (fact equal_refl)
  1113 
  1114 definition (in term_syntax)
  1115   bagify :: "('a\<Colon>typerep \<times> nat) list \<times> (unit \<Rightarrow> Code_Evaluation.term)
  1116     \<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where
  1117   [code_unfold]: "bagify xs = Code_Evaluation.valtermify Bag {\<cdot>} xs"
  1118 
  1119 notation fcomp (infixl "\<circ>>" 60)
  1120 notation scomp (infixl "\<circ>\<rightarrow>" 60)
  1121 
  1122 instantiation multiset :: (random) random
  1123 begin
  1124 
  1125 definition
  1126   "Quickcheck.random i = Quickcheck.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (bagify xs))"
  1127 
  1128 instance ..
  1129 
  1130 end
  1131 
  1132 no_notation fcomp (infixl "\<circ>>" 60)
  1133 no_notation scomp (infixl "\<circ>\<rightarrow>" 60)
  1134 
  1135 hide_const (open) bagify
  1136 
  1137 
  1138 subsection {* The multiset order *}
  1139 
  1140 subsubsection {* Well-foundedness *}
  1141 
  1142 definition mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
  1143   "mult1 r = {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
  1144       (\<forall>b. b :# K --> (b, a) \<in> r)}"
  1145 
  1146 definition mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
  1147   "mult r = (mult1 r)\<^sup>+"
  1148 
  1149 lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
  1150 by (simp add: mult1_def)
  1151 
  1152 lemma less_add: "(N, M0 + {#a#}) \<in> mult1 r ==>
  1153     (\<exists>M. (M, M0) \<in> mult1 r \<and> N = M + {#a#}) \<or>
  1154     (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K)"
  1155   (is "_ \<Longrightarrow> ?case1 (mult1 r) \<or> ?case2")
  1156 proof (unfold mult1_def)
  1157   let ?r = "\<lambda>K a. \<forall>b. b :# K --> (b, a) \<in> r"
  1158   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
  1159   let ?case1 = "?case1 {(N, M). ?R N M}"
  1160 
  1161   assume "(N, M0 + {#a#}) \<in> {(N, M). ?R N M}"
  1162   then have "\<exists>a' M0' K.
  1163       M0 + {#a#} = M0' + {#a'#} \<and> N = M0' + K \<and> ?r K a'" by simp
  1164   then show "?case1 \<or> ?case2"
  1165   proof (elim exE conjE)
  1166     fix a' M0' K
  1167     assume N: "N = M0' + K" and r: "?r K a'"
  1168     assume "M0 + {#a#} = M0' + {#a'#}"
  1169     then have "M0 = M0' \<and> a = a' \<or>
  1170         (\<exists>K'. M0 = K' + {#a'#} \<and> M0' = K' + {#a#})"
  1171       by (simp only: add_eq_conv_ex)
  1172     then show ?thesis
  1173     proof (elim disjE conjE exE)
  1174       assume "M0 = M0'" "a = a'"
  1175       with N r have "?r K a \<and> N = M0 + K" by simp
  1176       then have ?case2 .. then show ?thesis ..
  1177     next
  1178       fix K'
  1179       assume "M0' = K' + {#a#}"
  1180       with N have n: "N = K' + K + {#a#}" by (simp add: add_ac)
  1181 
  1182       assume "M0 = K' + {#a'#}"
  1183       with r have "?R (K' + K) M0" by blast
  1184       with n have ?case1 by simp then show ?thesis ..
  1185     qed
  1186   qed
  1187 qed
  1188 
  1189 lemma all_accessible: "wf r ==> \<forall>M. M \<in> acc (mult1 r)"
  1190 proof
  1191   let ?R = "mult1 r"
  1192   let ?W = "acc ?R"
  1193   {
  1194     fix M M0 a
  1195     assume M0: "M0 \<in> ?W"
  1196       and wf_hyp: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
  1197       and acc_hyp: "\<forall>M. (M, M0) \<in> ?R --> M + {#a#} \<in> ?W"
  1198     have "M0 + {#a#} \<in> ?W"
  1199     proof (rule accI [of "M0 + {#a#}"])
  1200       fix N
  1201       assume "(N, M0 + {#a#}) \<in> ?R"
  1202       then have "((\<exists>M. (M, M0) \<in> ?R \<and> N = M + {#a#}) \<or>
  1203           (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K))"
  1204         by (rule less_add)
  1205       then show "N \<in> ?W"
  1206       proof (elim exE disjE conjE)
  1207         fix M assume "(M, M0) \<in> ?R" and N: "N = M + {#a#}"
  1208         from acc_hyp have "(M, M0) \<in> ?R --> M + {#a#} \<in> ?W" ..
  1209         from this and `(M, M0) \<in> ?R` have "M + {#a#} \<in> ?W" ..
  1210         then show "N \<in> ?W" by (simp only: N)
  1211       next
  1212         fix K
  1213         assume N: "N = M0 + K"
  1214         assume "\<forall>b. b :# K --> (b, a) \<in> r"
  1215         then have "M0 + K \<in> ?W"
  1216         proof (induct K)
  1217           case empty
  1218           from M0 show "M0 + {#} \<in> ?W" by simp
  1219         next
  1220           case (add K x)
  1221           from add.prems have "(x, a) \<in> r" by simp
  1222           with wf_hyp have "\<forall>M \<in> ?W. M + {#x#} \<in> ?W" by blast
  1223           moreover from add have "M0 + K \<in> ?W" by simp
  1224           ultimately have "(M0 + K) + {#x#} \<in> ?W" ..
  1225           then show "M0 + (K + {#x#}) \<in> ?W" by (simp only: add_assoc)
  1226         qed
  1227         then show "N \<in> ?W" by (simp only: N)
  1228       qed
  1229     qed
  1230   } note tedious_reasoning = this
  1231 
  1232   assume wf: "wf r"
  1233   fix M
  1234   show "M \<in> ?W"
  1235   proof (induct M)
  1236     show "{#} \<in> ?W"
  1237     proof (rule accI)
  1238       fix b assume "(b, {#}) \<in> ?R"
  1239       with not_less_empty show "b \<in> ?W" by contradiction
  1240     qed
  1241 
  1242     fix M a assume "M \<in> ?W"
  1243     from wf have "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
  1244     proof induct
  1245       fix a
  1246       assume r: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
  1247       show "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
  1248       proof
  1249         fix M assume "M \<in> ?W"
  1250         then show "M + {#a#} \<in> ?W"
  1251           by (rule acc_induct) (rule tedious_reasoning [OF _ r])
  1252       qed
  1253     qed
  1254     from this and `M \<in> ?W` show "M + {#a#} \<in> ?W" ..
  1255   qed
  1256 qed
  1257 
  1258 theorem wf_mult1: "wf r ==> wf (mult1 r)"
  1259 by (rule acc_wfI) (rule all_accessible)
  1260 
  1261 theorem wf_mult: "wf r ==> wf (mult r)"
  1262 unfolding mult_def by (rule wf_trancl) (rule wf_mult1)
  1263 
  1264 
  1265 subsubsection {* Closure-free presentation *}
  1266 
  1267 text {* One direction. *}
  1268 
  1269 lemma mult_implies_one_step:
  1270   "trans r ==> (M, N) \<in> mult r ==>
  1271     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
  1272     (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r)"
  1273 apply (unfold mult_def mult1_def set_of_def)
  1274 apply (erule converse_trancl_induct, clarify)
  1275  apply (rule_tac x = M0 in exI, simp, clarify)
  1276 apply (case_tac "a :# K")
  1277  apply (rule_tac x = I in exI)
  1278  apply (simp (no_asm))
  1279  apply (rule_tac x = "(K - {#a#}) + Ka" in exI)
  1280  apply (simp (no_asm_simp) add: add_assoc [symmetric])
  1281  apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong)
  1282  apply (simp add: diff_union_single_conv)
  1283  apply (simp (no_asm_use) add: trans_def)
  1284  apply blast
  1285 apply (subgoal_tac "a :# I")
  1286  apply (rule_tac x = "I - {#a#}" in exI)
  1287  apply (rule_tac x = "J + {#a#}" in exI)
  1288  apply (rule_tac x = "K + Ka" in exI)
  1289  apply (rule conjI)
  1290   apply (simp add: multiset_eq_iff split: nat_diff_split)
  1291  apply (rule conjI)
  1292   apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong, simp)
  1293   apply (simp add: multiset_eq_iff split: nat_diff_split)
  1294  apply (simp (no_asm_use) add: trans_def)
  1295  apply blast
  1296 apply (subgoal_tac "a :# (M0 + {#a#})")
  1297  apply simp
  1298 apply (simp (no_asm))
  1299 done
  1300 
  1301 lemma one_step_implies_mult_aux:
  1302   "trans r ==>
  1303     \<forall>I J K. (size J = n \<and> J \<noteq> {#} \<and> (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r))
  1304       --> (I + K, I + J) \<in> mult r"
  1305 apply (induct_tac n, auto)
  1306 apply (frule size_eq_Suc_imp_eq_union, clarify)
  1307 apply (rename_tac "J'", simp)
  1308 apply (erule notE, auto)
  1309 apply (case_tac "J' = {#}")
  1310  apply (simp add: mult_def)
  1311  apply (rule r_into_trancl)
  1312  apply (simp add: mult1_def set_of_def, blast)
  1313 txt {* Now we know @{term "J' \<noteq> {#}"}. *}
  1314 apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition)
  1315 apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp)
  1316 apply (erule ssubst)
  1317 apply (simp add: Ball_def, auto)
  1318 apply (subgoal_tac
  1319   "((I + {# x :# K. (x, a) \<in> r #}) + {# x :# K. (x, a) \<notin> r #},
  1320     (I + {# x :# K. (x, a) \<in> r #}) + J') \<in> mult r")
  1321  prefer 2
  1322  apply force
  1323 apply (simp (no_asm_use) add: add_assoc [symmetric] mult_def)
  1324 apply (erule trancl_trans)
  1325 apply (rule r_into_trancl)
  1326 apply (simp add: mult1_def set_of_def)
  1327 apply (rule_tac x = a in exI)
  1328 apply (rule_tac x = "I + J'" in exI)
  1329 apply (simp add: add_ac)
  1330 done
  1331 
  1332 lemma one_step_implies_mult:
  1333   "trans r ==> J \<noteq> {#} ==> \<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r
  1334     ==> (I + K, I + J) \<in> mult r"
  1335 using one_step_implies_mult_aux by blast
  1336 
  1337 
  1338 subsubsection {* Partial-order properties *}
  1339 
  1340 definition less_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<#" 50) where
  1341   "M' <# M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
  1342 
  1343 definition le_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<=#" 50) where
  1344   "M' <=# M \<longleftrightarrow> M' <# M \<or> M' = M"
  1345 
  1346 notation (xsymbols) less_multiset (infix "\<subset>#" 50)
  1347 notation (xsymbols) le_multiset (infix "\<subseteq>#" 50)
  1348 
  1349 interpretation multiset_order: order le_multiset less_multiset
  1350 proof -
  1351   have irrefl: "\<And>M :: 'a multiset. \<not> M \<subset># M"
  1352   proof
  1353     fix M :: "'a multiset"
  1354     assume "M \<subset># M"
  1355     then have MM: "(M, M) \<in> mult {(x, y). x < y}" by (simp add: less_multiset_def)
  1356     have "trans {(x'::'a, x). x' < x}"
  1357       by (rule transI) simp
  1358     moreover note MM
  1359     ultimately have "\<exists>I J K. M = I + J \<and> M = I + K
  1360       \<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})"
  1361       by (rule mult_implies_one_step)
  1362     then obtain I J K where "M = I + J" and "M = I + K"
  1363       and "J \<noteq> {#}" and "(\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})" by blast
  1364     then have aux1: "K \<noteq> {#}" and aux2: "\<forall>k\<in>set_of K. \<exists>j\<in>set_of K. k < j" by auto
  1365     have "finite (set_of K)" by simp
  1366     moreover note aux2
  1367     ultimately have "set_of K = {}"
  1368       by (induct rule: finite_induct) (auto intro: order_less_trans)
  1369     with aux1 show False by simp
  1370   qed
  1371   have trans: "\<And>K M N :: 'a multiset. K \<subset># M \<Longrightarrow> M \<subset># N \<Longrightarrow> K \<subset># N"
  1372     unfolding less_multiset_def mult_def by (blast intro: trancl_trans)
  1373   show "class.order (le_multiset :: 'a multiset \<Rightarrow> _) less_multiset" proof
  1374   qed (auto simp add: le_multiset_def irrefl dest: trans)
  1375 qed
  1376 
  1377 lemma mult_less_irrefl [elim!]:
  1378   "M \<subset># (M::'a::order multiset) ==> R"
  1379   by (simp add: multiset_order.less_irrefl)
  1380 
  1381 
  1382 subsubsection {* Monotonicity of multiset union *}
  1383 
  1384 lemma mult1_union:
  1385   "(B, D) \<in> mult1 r ==> (C + B, C + D) \<in> mult1 r"
  1386 apply (unfold mult1_def)
  1387 apply auto
  1388 apply (rule_tac x = a in exI)
  1389 apply (rule_tac x = "C + M0" in exI)
  1390 apply (simp add: add_assoc)
  1391 done
  1392 
  1393 lemma union_less_mono2: "B \<subset># D ==> C + B \<subset># C + (D::'a::order multiset)"
  1394 apply (unfold less_multiset_def mult_def)
  1395 apply (erule trancl_induct)
  1396  apply (blast intro: mult1_union)
  1397 apply (blast intro: mult1_union trancl_trans)
  1398 done
  1399 
  1400 lemma union_less_mono1: "B \<subset># D ==> B + C \<subset># D + (C::'a::order multiset)"
  1401 apply (subst add_commute [of B C])
  1402 apply (subst add_commute [of D C])
  1403 apply (erule union_less_mono2)
  1404 done
  1405 
  1406 lemma union_less_mono:
  1407   "A \<subset># C ==> B \<subset># D ==> A + B \<subset># C + (D::'a::order multiset)"
  1408   by (blast intro!: union_less_mono1 union_less_mono2 multiset_order.less_trans)
  1409 
  1410 interpretation multiset_order: ordered_ab_semigroup_add plus le_multiset less_multiset
  1411 proof
  1412 qed (auto simp add: le_multiset_def intro: union_less_mono2)
  1413 
  1414 
  1415 subsection {* The fold combinator *}
  1416 
  1417 text {*
  1418   The intended behaviour is
  1419   @{text "fold_mset f z {#x\<^isub>1, ..., x\<^isub>n#} = f x\<^isub>1 (\<dots> (f x\<^isub>n z)\<dots>)"}
  1420   if @{text f} is associative-commutative. 
  1421 *}
  1422 
  1423 text {*
  1424   The graph of @{text "fold_mset"}, @{text "z"}: the start element,
  1425   @{text "f"}: folding function, @{text "A"}: the multiset, @{text
  1426   "y"}: the result.
  1427 *}
  1428 inductive 
  1429   fold_msetG :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b \<Rightarrow> bool" 
  1430   for f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" 
  1431   and z :: 'b
  1432 where
  1433   emptyI [intro]:  "fold_msetG f z {#} z"
  1434 | insertI [intro]: "fold_msetG f z A y \<Longrightarrow> fold_msetG f z (A + {#x#}) (f x y)"
  1435 
  1436 inductive_cases empty_fold_msetGE [elim!]: "fold_msetG f z {#} x"
  1437 inductive_cases insert_fold_msetGE: "fold_msetG f z (A + {#}) y" 
  1438 
  1439 definition
  1440   fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b" where
  1441   "fold_mset f z A = (THE x. fold_msetG f z A x)"
  1442 
  1443 lemma Diff1_fold_msetG:
  1444   "fold_msetG f z (A - {#x#}) y \<Longrightarrow> x \<in># A \<Longrightarrow> fold_msetG f z A (f x y)"
  1445 apply (frule_tac x = x in fold_msetG.insertI)
  1446 apply auto
  1447 done
  1448 
  1449 lemma fold_msetG_nonempty: "\<exists>x. fold_msetG f z A x"
  1450 apply (induct A)
  1451  apply blast
  1452 apply clarsimp
  1453 apply (drule_tac x = x in fold_msetG.insertI)
  1454 apply auto
  1455 done
  1456 
  1457 lemma fold_mset_empty[simp]: "fold_mset f z {#} = z"
  1458 unfolding fold_mset_def by blast
  1459 
  1460 context fun_left_comm
  1461 begin
  1462 
  1463 lemma fold_msetG_determ:
  1464   "fold_msetG f z A x \<Longrightarrow> fold_msetG f z A y \<Longrightarrow> y = x"
  1465 proof (induct arbitrary: x y z rule: full_multiset_induct)
  1466   case (less M x\<^isub>1 x\<^isub>2 Z)
  1467   have IH: "\<forall>A. A < M \<longrightarrow> 
  1468     (\<forall>x x' x''. fold_msetG f x'' A x \<longrightarrow> fold_msetG f x'' A x'
  1469                \<longrightarrow> x' = x)" by fact
  1470   have Mfoldx\<^isub>1: "fold_msetG f Z M x\<^isub>1" and Mfoldx\<^isub>2: "fold_msetG f Z M x\<^isub>2" by fact+
  1471   show ?case
  1472   proof (rule fold_msetG.cases [OF Mfoldx\<^isub>1])
  1473     assume "M = {#}" and "x\<^isub>1 = Z"
  1474     then show ?case using Mfoldx\<^isub>2 by auto 
  1475   next
  1476     fix B b u
  1477     assume "M = B + {#b#}" and "x\<^isub>1 = f b u" and Bu: "fold_msetG f Z B u"
  1478     then have MBb: "M = B + {#b#}" and x\<^isub>1: "x\<^isub>1 = f b u" by auto
  1479     show ?case
  1480     proof (rule fold_msetG.cases [OF Mfoldx\<^isub>2])
  1481       assume "M = {#}" "x\<^isub>2 = Z"
  1482       then show ?case using Mfoldx\<^isub>1 by auto
  1483     next
  1484       fix C c v
  1485       assume "M = C + {#c#}" and "x\<^isub>2 = f c v" and Cv: "fold_msetG f Z C v"
  1486       then have MCc: "M = C + {#c#}" and x\<^isub>2: "x\<^isub>2 = f c v" by auto
  1487       then have CsubM: "C < M" by simp
  1488       from MBb have BsubM: "B < M" by simp
  1489       show ?case
  1490       proof cases
  1491         assume "b=c"
  1492         then moreover have "B = C" using MBb MCc by auto
  1493         ultimately show ?thesis using Bu Cv x\<^isub>1 x\<^isub>2 CsubM IH by auto
  1494       next
  1495         assume diff: "b \<noteq> c"
  1496         let ?D = "B - {#c#}"
  1497         have cinB: "c \<in># B" and binC: "b \<in># C" using MBb MCc diff
  1498           by (auto intro: insert_noteq_member dest: sym)
  1499         have "B - {#c#} < B" using cinB by (rule mset_less_diff_self)
  1500         then have DsubM: "?D < M" using BsubM by (blast intro: order_less_trans)
  1501         from MBb MCc have "B + {#b#} = C + {#c#}" by blast
  1502         then have [simp]: "B + {#b#} - {#c#} = C"
  1503           using MBb MCc binC cinB by auto
  1504         have B: "B = ?D + {#c#}" and C: "C = ?D + {#b#}"
  1505           using MBb MCc diff binC cinB
  1506           by (auto simp: multiset_add_sub_el_shuffle)
  1507         then obtain d where Dfoldd: "fold_msetG f Z ?D d"
  1508           using fold_msetG_nonempty by iprover
  1509         then have "fold_msetG f Z B (f c d)" using cinB
  1510           by (rule Diff1_fold_msetG)
  1511         then have "f c d = u" using IH BsubM Bu by blast
  1512         moreover 
  1513         have "fold_msetG f Z C (f b d)" using binC cinB diff Dfoldd
  1514           by (auto simp: multiset_add_sub_el_shuffle 
  1515             dest: fold_msetG.insertI [where x=b])
  1516         then have "f b d = v" using IH CsubM Cv by blast
  1517         ultimately show ?thesis using x\<^isub>1 x\<^isub>2
  1518           by (auto simp: fun_left_comm)
  1519       qed
  1520     qed
  1521   qed
  1522 qed
  1523         
  1524 lemma fold_mset_insert_aux:
  1525   "(fold_msetG f z (A + {#x#}) v) =
  1526     (\<exists>y. fold_msetG f z A y \<and> v = f x y)"
  1527 apply (rule iffI)
  1528  prefer 2
  1529  apply blast
  1530 apply (rule_tac A=A and f=f in fold_msetG_nonempty [THEN exE, standard])
  1531 apply (blast intro: fold_msetG_determ)
  1532 done
  1533 
  1534 lemma fold_mset_equality: "fold_msetG f z A y \<Longrightarrow> fold_mset f z A = y"
  1535 unfolding fold_mset_def by (blast intro: fold_msetG_determ)
  1536 
  1537 lemma fold_mset_insert:
  1538   "fold_mset f z (A + {#x#}) = f x (fold_mset f z A)"
  1539 apply (simp add: fold_mset_def fold_mset_insert_aux)
  1540 apply (rule the_equality)
  1541  apply (auto cong add: conj_cong 
  1542      simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty)
  1543 done
  1544 
  1545 lemma fold_mset_commute: "f x (fold_mset f z A) = fold_mset f (f x z) A"
  1546 by (induct A) (auto simp: fold_mset_insert fun_left_comm [of x])
  1547 
  1548 lemma fold_mset_single [simp]: "fold_mset f z {#x#} = f x z"
  1549 using fold_mset_insert [of z "{#}"] by simp
  1550 
  1551 lemma fold_mset_union [simp]:
  1552   "fold_mset f z (A+B) = fold_mset f (fold_mset f z A) B"
  1553 proof (induct A)
  1554   case empty then show ?case by simp
  1555 next
  1556   case (add A x)
  1557   have "A + {#x#} + B = (A+B) + {#x#}" by (simp add: add_ac)
  1558   then have "fold_mset f z (A + {#x#} + B) = f x (fold_mset f z (A + B))" 
  1559     by (simp add: fold_mset_insert)
  1560   also have "\<dots> = fold_mset f (fold_mset f z (A + {#x#})) B"
  1561     by (simp add: fold_mset_commute[of x,symmetric] add fold_mset_insert)
  1562   finally show ?case .
  1563 qed
  1564 
  1565 lemma fold_mset_fusion:
  1566   assumes "fun_left_comm g"
  1567   shows "(\<And>x y. h (g x y) = f x (h y)) \<Longrightarrow> h (fold_mset g w A) = fold_mset f (h w) A" (is "PROP ?P")
  1568 proof -
  1569   interpret fun_left_comm g by (fact assms)
  1570   show "PROP ?P" by (induct A) auto
  1571 qed
  1572 
  1573 lemma fold_mset_rec:
  1574   assumes "a \<in># A" 
  1575   shows "fold_mset f z A = f a (fold_mset f z (A - {#a#}))"
  1576 proof -
  1577   from assms obtain A' where "A = A' + {#a#}"
  1578     by (blast dest: multi_member_split)
  1579   then show ?thesis by simp
  1580 qed
  1581 
  1582 end
  1583 
  1584 text {*
  1585   A note on code generation: When defining some function containing a
  1586   subterm @{term"fold_mset F"}, code generation is not automatic. When
  1587   interpreting locale @{text left_commutative} with @{text F}, the
  1588   would be code thms for @{const fold_mset} become thms like
  1589   @{term"fold_mset F z {#} = z"} where @{text F} is not a pattern but
  1590   contains defined symbols, i.e.\ is not a code thm. Hence a separate
  1591   constant with its own code thms needs to be introduced for @{text
  1592   F}. See the image operator below.
  1593 *}
  1594 
  1595 
  1596 subsection {* Image *}
  1597 
  1598 definition image_mset :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
  1599   "image_mset f = fold_mset (op + o single o f) {#}"
  1600 
  1601 interpretation image_left_comm: fun_left_comm "op + o single o f"
  1602 proof qed (simp add: add_ac fun_eq_iff)
  1603 
  1604 lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
  1605 by (simp add: image_mset_def)
  1606 
  1607 lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}"
  1608 by (simp add: image_mset_def)
  1609 
  1610 lemma image_mset_insert:
  1611   "image_mset f (M + {#a#}) = image_mset f M + {#f a#}"
  1612 by (simp add: image_mset_def add_ac)
  1613 
  1614 lemma image_mset_union [simp]:
  1615   "image_mset f (M+N) = image_mset f M + image_mset f N"
  1616 apply (induct N)
  1617  apply simp
  1618 apply (simp add: add_assoc [symmetric] image_mset_insert)
  1619 done
  1620 
  1621 lemma size_image_mset [simp]: "size (image_mset f M) = size M"
  1622 by (induct M) simp_all
  1623 
  1624 lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \<longleftrightarrow> M = {#}"
  1625 by (cases M) auto
  1626 
  1627 syntax
  1628   "_comprehension1_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"
  1629       ("({#_/. _ :# _#})")
  1630 translations
  1631   "{#e. x:#M#}" == "CONST image_mset (%x. e) M"
  1632 
  1633 syntax
  1634   "_comprehension2_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"
  1635       ("({#_/ | _ :# _./ _#})")
  1636 translations
  1637   "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}"
  1638 
  1639 text {*
  1640   This allows to write not just filters like @{term "{#x:#M. x<c#}"}
  1641   but also images like @{term "{#x+x. x:#M #}"} and @{term [source]
  1642   "{#x+x|x:#M. x<c#}"}, where the latter is currently displayed as
  1643   @{term "{#x+x|x:#M. x<c#}"}.
  1644 *}
  1645 
  1646 enriched_type image_mset: image_mset proof -
  1647   fix f g 
  1648   show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)"
  1649   proof
  1650     fix A
  1651     show "(image_mset f \<circ> image_mset g) A = image_mset (f \<circ> g) A"
  1652       by (induct A) simp_all
  1653   qed
  1654 next
  1655   show "image_mset id = id"
  1656   proof
  1657     fix A
  1658     show "image_mset id A = id A"
  1659       by (induct A) simp_all
  1660   qed
  1661 qed
  1662 
  1663 
  1664 subsection {* Termination proofs with multiset orders *}
  1665 
  1666 lemma multi_member_skip: "x \<in># XS \<Longrightarrow> x \<in># {# y #} + XS"
  1667   and multi_member_this: "x \<in># {# x #} + XS"
  1668   and multi_member_last: "x \<in># {# x #}"
  1669   by auto
  1670 
  1671 definition "ms_strict = mult pair_less"
  1672 definition "ms_weak = ms_strict \<union> Id"
  1673 
  1674 lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)"
  1675 unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def
  1676 by (auto intro: wf_mult1 wf_trancl simp: mult_def)
  1677 
  1678 lemma smsI:
  1679   "(set_of A, set_of B) \<in> max_strict \<Longrightarrow> (Z + A, Z + B) \<in> ms_strict"
  1680   unfolding ms_strict_def
  1681 by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases)
  1682 
  1683 lemma wmsI:
  1684   "(set_of A, set_of B) \<in> max_strict \<or> A = {#} \<and> B = {#}
  1685   \<Longrightarrow> (Z + A, Z + B) \<in> ms_weak"
  1686 unfolding ms_weak_def ms_strict_def
  1687 by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult)
  1688 
  1689 inductive pw_leq
  1690 where
  1691   pw_leq_empty: "pw_leq {#} {#}"
  1692 | pw_leq_step:  "\<lbrakk>(x,y) \<in> pair_leq; pw_leq X Y \<rbrakk> \<Longrightarrow> pw_leq ({#x#} + X) ({#y#} + Y)"
  1693 
  1694 lemma pw_leq_lstep:
  1695   "(x, y) \<in> pair_leq \<Longrightarrow> pw_leq {#x#} {#y#}"
  1696 by (drule pw_leq_step) (rule pw_leq_empty, simp)
  1697 
  1698 lemma pw_leq_split:
  1699   assumes "pw_leq X Y"
  1700   shows "\<exists>A B Z. X = A + Z \<and> Y = B + Z \<and> ((set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
  1701   using assms
  1702 proof (induct)
  1703   case pw_leq_empty thus ?case by auto
  1704 next
  1705   case (pw_leq_step x y X Y)
  1706   then obtain A B Z where
  1707     [simp]: "X = A + Z" "Y = B + Z" 
  1708       and 1[simp]: "(set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#})" 
  1709     by auto
  1710   from pw_leq_step have "x = y \<or> (x, y) \<in> pair_less" 
  1711     unfolding pair_leq_def by auto
  1712   thus ?case
  1713   proof
  1714     assume [simp]: "x = y"
  1715     have
  1716       "{#x#} + X = A + ({#y#}+Z) 
  1717       \<and> {#y#} + Y = B + ({#y#}+Z)
  1718       \<and> ((set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
  1719       by (auto simp: add_ac)
  1720     thus ?case by (intro exI)
  1721   next
  1722     assume A: "(x, y) \<in> pair_less"
  1723     let ?A' = "{#x#} + A" and ?B' = "{#y#} + B"
  1724     have "{#x#} + X = ?A' + Z"
  1725       "{#y#} + Y = ?B' + Z"
  1726       by (auto simp add: add_ac)
  1727     moreover have 
  1728       "(set_of ?A', set_of ?B') \<in> max_strict"
  1729       using 1 A unfolding max_strict_def 
  1730       by (auto elim!: max_ext.cases)
  1731     ultimately show ?thesis by blast
  1732   qed
  1733 qed
  1734 
  1735 lemma 
  1736   assumes pwleq: "pw_leq Z Z'"
  1737   shows ms_strictI: "(set_of A, set_of B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_strict"
  1738   and   ms_weakI1:  "(set_of A, set_of B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_weak"
  1739   and   ms_weakI2:  "(Z + {#}, Z' + {#}) \<in> ms_weak"
  1740 proof -
  1741   from pw_leq_split[OF pwleq] 
  1742   obtain A' B' Z''
  1743     where [simp]: "Z = A' + Z''" "Z' = B' + Z''"
  1744     and mx_or_empty: "(set_of A', set_of B') \<in> max_strict \<or> (A' = {#} \<and> B' = {#})"
  1745     by blast
  1746   {
  1747     assume max: "(set_of A, set_of B) \<in> max_strict"
  1748     from mx_or_empty
  1749     have "(Z'' + (A + A'), Z'' + (B + B')) \<in> ms_strict"
  1750     proof
  1751       assume max': "(set_of A', set_of B') \<in> max_strict"
  1752       with max have "(set_of (A + A'), set_of (B + B')) \<in> max_strict"
  1753         by (auto simp: max_strict_def intro: max_ext_additive)
  1754       thus ?thesis by (rule smsI) 
  1755     next
  1756       assume [simp]: "A' = {#} \<and> B' = {#}"
  1757       show ?thesis by (rule smsI) (auto intro: max)
  1758     qed
  1759     thus "(Z + A, Z' + B) \<in> ms_strict" by (simp add:add_ac)
  1760     thus "(Z + A, Z' + B) \<in> ms_weak" by (simp add: ms_weak_def)
  1761   }
  1762   from mx_or_empty
  1763   have "(Z'' + A', Z'' + B') \<in> ms_weak" by (rule wmsI)
  1764   thus "(Z + {#}, Z' + {#}) \<in> ms_weak" by (simp add:add_ac)
  1765 qed
  1766 
  1767 lemma empty_neutral: "{#} + x = x" "x + {#} = x"
  1768 and nonempty_plus: "{# x #} + rs \<noteq> {#}"
  1769 and nonempty_single: "{# x #} \<noteq> {#}"
  1770 by auto
  1771 
  1772 setup {*
  1773 let
  1774   fun msetT T = Type (@{type_name multiset}, [T]);
  1775 
  1776   fun mk_mset T [] = Const (@{const_abbrev Mempty}, msetT T)
  1777     | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x
  1778     | mk_mset T (x :: xs) =
  1779           Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $
  1780                 mk_mset T [x] $ mk_mset T xs
  1781 
  1782   fun mset_member_tac m i =
  1783       (if m <= 0 then
  1784            rtac @{thm multi_member_this} i ORELSE rtac @{thm multi_member_last} i
  1785        else
  1786            rtac @{thm multi_member_skip} i THEN mset_member_tac (m - 1) i)
  1787 
  1788   val mset_nonempty_tac =
  1789       rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single}
  1790 
  1791   val regroup_munion_conv =
  1792       Function_Lib.regroup_conv @{const_abbrev Mempty} @{const_name plus}
  1793         (map (fn t => t RS eq_reflection) (@{thms add_ac} @ @{thms empty_neutral}))
  1794 
  1795   fun unfold_pwleq_tac i =
  1796     (rtac @{thm pw_leq_step} i THEN (fn st => unfold_pwleq_tac (i + 1) st))
  1797       ORELSE (rtac @{thm pw_leq_lstep} i)
  1798       ORELSE (rtac @{thm pw_leq_empty} i)
  1799 
  1800   val set_of_simps = [@{thm set_of_empty}, @{thm set_of_single}, @{thm set_of_union},
  1801                       @{thm Un_insert_left}, @{thm Un_empty_left}]
  1802 in
  1803   ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset 
  1804   {
  1805     msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv,
  1806     mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac,
  1807     mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_of_simps,
  1808     smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1},
  1809     reduction_pair= @{thm ms_reduction_pair}
  1810   })
  1811 end
  1812 *}
  1813 
  1814 
  1815 subsection {* Legacy theorem bindings *}
  1816 
  1817 lemmas multi_count_eq = multiset_eq_iff [symmetric]
  1818 
  1819 lemma union_commute: "M + N = N + (M::'a multiset)"
  1820   by (fact add_commute)
  1821 
  1822 lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))"
  1823   by (fact add_assoc)
  1824 
  1825 lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))"
  1826   by (fact add_left_commute)
  1827 
  1828 lemmas union_ac = union_assoc union_commute union_lcomm
  1829 
  1830 lemma union_right_cancel: "M + K = N + K \<longleftrightarrow> M = (N::'a multiset)"
  1831   by (fact add_right_cancel)
  1832 
  1833 lemma union_left_cancel: "K + M = K + N \<longleftrightarrow> M = (N::'a multiset)"
  1834   by (fact add_left_cancel)
  1835 
  1836 lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y"
  1837   by (fact add_imp_eq)
  1838 
  1839 lemma mset_less_trans: "(M::'a multiset) < K \<Longrightarrow> K < N \<Longrightarrow> M < N"
  1840   by (fact order_less_trans)
  1841 
  1842 lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A"
  1843   by (fact inf.commute)
  1844 
  1845 lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C"
  1846   by (fact inf.assoc [symmetric])
  1847 
  1848 lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)"
  1849   by (fact inf.left_commute)
  1850 
  1851 lemmas multiset_inter_ac =
  1852   multiset_inter_commute
  1853   multiset_inter_assoc
  1854   multiset_inter_left_commute
  1855 
  1856 lemma mult_less_not_refl:
  1857   "\<not> M \<subset># (M::'a::order multiset)"
  1858   by (fact multiset_order.less_irrefl)
  1859 
  1860 lemma mult_less_trans:
  1861   "K \<subset># M ==> M \<subset># N ==> K \<subset># (N::'a::order multiset)"
  1862   by (fact multiset_order.less_trans)
  1863     
  1864 lemma mult_less_not_sym:
  1865   "M \<subset># N ==> \<not> N \<subset># (M::'a::order multiset)"
  1866   by (fact multiset_order.less_not_sym)
  1867 
  1868 lemma mult_less_asym:
  1869   "M \<subset># N ==> (\<not> P ==> N \<subset># (M::'a::order multiset)) ==> P"
  1870   by (fact multiset_order.less_asym)
  1871 
  1872 ML {*
  1873 fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T]))
  1874                       (Const _ $ t') =
  1875     let
  1876       val (maybe_opt, ps) =
  1877         Nitpick_Model.dest_plain_fun t' ||> op ~~
  1878         ||> map (apsnd (snd o HOLogic.dest_number))
  1879       fun elems_for t =
  1880         case AList.lookup (op =) ps t of
  1881           SOME n => replicate n t
  1882         | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]
  1883     in
  1884       case maps elems_for (all_values elem_T) @
  1885            (if maybe_opt then [Const (Nitpick_Model.unrep (), elem_T)]
  1886             else []) of
  1887         [] => Const (@{const_name zero_class.zero}, T)
  1888       | ts => foldl1 (fn (t1, t2) =>
  1889                          Const (@{const_name plus_class.plus}, T --> T --> T)
  1890                          $ t1 $ t2)
  1891                      (map (curry (op $) (Const (@{const_name single},
  1892                                                 elem_T --> T))) ts)
  1893     end
  1894   | multiset_postproc _ _ _ _ t = t
  1895 *}
  1896 
  1897 declaration {*
  1898 Nitpick_Model.register_term_postprocessor @{typ "'a multiset"}
  1899     multiset_postproc
  1900 *}
  1901 
  1902 end