| author | webertj | 
| Tue, 01 Aug 2006 14:58:43 +0200 | |
| changeset 20276 | d94dc40673b1 | 
| parent 19582 | a669c98b9c24 | 
| child 20503 | 503ac4c5ef91 | 
| permissions | -rw-r--r-- | 
| 10249 | 1 | (* Title: HOL/Library/Multiset.thy | 
| 2 | ID: $Id$ | |
| 15072 | 3 | Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker | 
| 10249 | 4 | *) | 
| 5 | ||
| 14706 | 6 | header {* Multisets *}
 | 
| 10249 | 7 | |
| 15131 | 8 | theory Multiset | 
| 19564 
d3e2f532459a
First usable version of the new function definition package (HOL/function_packake/...).
 krauss parents: 
19363diff
changeset | 9 | imports Main | 
| 15131 | 10 | begin | 
| 10249 | 11 | |
| 12 | subsection {* The type of multisets *}
 | |
| 13 | ||
| 14 | typedef 'a multiset = "{f::'a => nat. finite {x . 0 < f x}}"
 | |
| 15 | proof | |
| 11464 | 16 | show "(\<lambda>x. 0::nat) \<in> ?multiset" by simp | 
| 10249 | 17 | qed | 
| 18 | ||
| 19 | lemmas multiset_typedef [simp] = | |
| 10277 | 20 | Abs_multiset_inverse Rep_multiset_inverse Rep_multiset | 
| 21 | and [simp] = Rep_multiset_inject [symmetric] | |
| 10249 | 22 | |
| 19086 | 23 | definition | 
| 10249 | 24 |   Mempty :: "'a multiset"    ("{#}")
 | 
| 19086 | 25 |   "{#} = Abs_multiset (\<lambda>a. 0)"
 | 
| 10249 | 26 | |
| 27 |   single :: "'a => 'a multiset"    ("{#_#}")
 | |
| 19086 | 28 |   "{#a#} = Abs_multiset (\<lambda>b. if b = a then 1 else 0)"
 | 
| 10249 | 29 | |
| 30 | count :: "'a multiset => 'a => nat" | |
| 19086 | 31 | "count = Rep_multiset" | 
| 10249 | 32 | |
| 33 |   MCollect :: "'a multiset => ('a => bool) => 'a multiset"
 | |
| 19086 | 34 | "MCollect M P = Abs_multiset (\<lambda>x. if P x then Rep_multiset M x else 0)" | 
| 35 | ||
| 19363 | 36 | abbreviation | 
| 19086 | 37 |   Melem :: "'a => 'a multiset => bool"    ("(_/ :# _)" [50, 51] 50)
 | 
| 19363 | 38 | "a :# M == 0 < count M a" | 
| 10249 | 39 | |
| 40 | syntax | |
| 41 |   "_MCollect" :: "pttrn => 'a multiset => bool => 'a multiset"    ("(1{# _ : _./ _#})")
 | |
| 42 | translations | |
| 11464 | 43 |   "{#x:M. P#}" == "MCollect M (\<lambda>x. P)"
 | 
| 10249 | 44 | |
| 19086 | 45 | definition | 
| 10249 | 46 | set_of :: "'a multiset => 'a set" | 
| 19086 | 47 |   "set_of M = {x. x :# M}"
 | 
| 10249 | 48 | |
| 14691 | 49 | instance multiset :: (type) "{plus, minus, zero}" ..
 | 
| 10249 | 50 | |
| 51 | defs (overloaded) | |
| 11464 | 52 | union_def: "M + N == Abs_multiset (\<lambda>a. Rep_multiset M a + Rep_multiset N a)" | 
| 53 | diff_def: "M - N == Abs_multiset (\<lambda>a. Rep_multiset M a - Rep_multiset N a)" | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11655diff
changeset | 54 |   Zero_multiset_def [simp]: "0 == {#}"
 | 
| 10249 | 55 | size_def: "size M == setsum (count M) (set_of M)" | 
| 56 | ||
| 19086 | 57 | definition | 
| 58 | multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) | |
| 59 | "multiset_inter A B = A - (A - B)" | |
| 15869 | 60 | |
| 10249 | 61 | |
| 62 | text {*
 | |
| 63 |  \medskip Preservation of the representing set @{term multiset}.
 | |
| 64 | *} | |
| 65 | ||
| 11464 | 66 | lemma const0_in_multiset [simp]: "(\<lambda>a. 0) \<in> multiset" | 
| 17161 | 67 | by (simp add: multiset_def) | 
| 10249 | 68 | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11655diff
changeset | 69 | lemma only1_in_multiset [simp]: "(\<lambda>b. if b = a then 1 else 0) \<in> multiset" | 
| 17161 | 70 | by (simp add: multiset_def) | 
| 10249 | 71 | |
| 72 | lemma union_preserves_multiset [simp]: | |
| 11464 | 73 | "M \<in> multiset ==> N \<in> multiset ==> (\<lambda>a. M a + N a) \<in> multiset" | 
| 17161 | 74 | apply (simp add: multiset_def) | 
| 75 | apply (drule (1) finite_UnI) | |
| 10249 | 76 | apply (simp del: finite_Un add: Un_def) | 
| 77 | done | |
| 78 | ||
| 79 | lemma diff_preserves_multiset [simp]: | |
| 11464 | 80 | "M \<in> multiset ==> (\<lambda>a. M a - N a) \<in> multiset" | 
| 17161 | 81 | apply (simp add: multiset_def) | 
| 10249 | 82 | apply (rule finite_subset) | 
| 17161 | 83 | apply auto | 
| 10249 | 84 | done | 
| 85 | ||
| 86 | ||
| 87 | subsection {* Algebraic properties of multisets *}
 | |
| 88 | ||
| 89 | subsubsection {* Union *}
 | |
| 90 | ||
| 17161 | 91 | lemma union_empty [simp]: "M + {#} = M \<and> {#} + M = M"
 | 
| 92 | by (simp add: union_def Mempty_def) | |
| 10249 | 93 | |
| 17161 | 94 | lemma union_commute: "M + N = N + (M::'a multiset)" | 
| 95 | by (simp add: union_def add_ac) | |
| 96 | ||
| 97 | lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" | |
| 98 | by (simp add: union_def add_ac) | |
| 10249 | 99 | |
| 17161 | 100 | lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" | 
| 101 | proof - | |
| 102 | have "M + (N + K) = (N + K) + M" | |
| 103 | by (rule union_commute) | |
| 104 | also have "\<dots> = N + (K + M)" | |
| 105 | by (rule union_assoc) | |
| 106 | also have "K + M = M + K" | |
| 107 | by (rule union_commute) | |
| 108 | finally show ?thesis . | |
| 109 | qed | |
| 10249 | 110 | |
| 17161 | 111 | lemmas union_ac = union_assoc union_commute union_lcomm | 
| 10249 | 112 | |
| 14738 | 113 | instance multiset :: (type) comm_monoid_add | 
| 17200 | 114 | proof | 
| 14722 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 115 | fix a b c :: "'a multiset" | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 116 | show "(a + b) + c = a + (b + c)" by (rule union_assoc) | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 117 | show "a + b = b + a" by (rule union_commute) | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 118 | show "0 + a = a" by simp | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 119 | qed | 
| 10277 | 120 | |
| 10249 | 121 | |
| 122 | subsubsection {* Difference *}
 | |
| 123 | ||
| 17161 | 124 | lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
 | 
| 125 | by (simp add: Mempty_def diff_def) | |
| 10249 | 126 | |
| 17161 | 127 | lemma diff_union_inverse2 [simp]: "M + {#a#} - {#a#} = M"
 | 
| 128 | by (simp add: union_def diff_def) | |
| 10249 | 129 | |
| 130 | ||
| 131 | subsubsection {* Count of elements *}
 | |
| 132 | ||
| 17161 | 133 | lemma count_empty [simp]: "count {#} a = 0"
 | 
| 134 | by (simp add: count_def Mempty_def) | |
| 10249 | 135 | |
| 17161 | 136 | lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)"
 | 
| 137 | by (simp add: count_def single_def) | |
| 10249 | 138 | |
| 17161 | 139 | lemma count_union [simp]: "count (M + N) a = count M a + count N a" | 
| 140 | by (simp add: count_def union_def) | |
| 10249 | 141 | |
| 17161 | 142 | lemma count_diff [simp]: "count (M - N) a = count M a - count N a" | 
| 143 | by (simp add: count_def diff_def) | |
| 10249 | 144 | |
| 145 | ||
| 146 | subsubsection {* Set of elements *}
 | |
| 147 | ||
| 17161 | 148 | lemma set_of_empty [simp]: "set_of {#} = {}"
 | 
| 149 | by (simp add: set_of_def) | |
| 10249 | 150 | |
| 17161 | 151 | lemma set_of_single [simp]: "set_of {#b#} = {b}"
 | 
| 152 | by (simp add: set_of_def) | |
| 10249 | 153 | |
| 17161 | 154 | lemma set_of_union [simp]: "set_of (M + N) = set_of M \<union> set_of N" | 
| 155 | by (auto simp add: set_of_def) | |
| 10249 | 156 | |
| 17161 | 157 | lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})"
 | 
| 158 | by (auto simp add: set_of_def Mempty_def count_def expand_fun_eq) | |
| 10249 | 159 | |
| 17161 | 160 | lemma mem_set_of_iff [simp]: "(x \<in> set_of M) = (x :# M)" | 
| 161 | by (auto simp add: set_of_def) | |
| 10249 | 162 | |
| 163 | ||
| 164 | subsubsection {* Size *}
 | |
| 165 | ||
| 17161 | 166 | lemma size_empty [simp]: "size {#} = 0"
 | 
| 167 | by (simp add: size_def) | |
| 10249 | 168 | |
| 17161 | 169 | lemma size_single [simp]: "size {#b#} = 1"
 | 
| 170 | by (simp add: size_def) | |
| 10249 | 171 | |
| 17161 | 172 | lemma finite_set_of [iff]: "finite (set_of M)" | 
| 173 | using Rep_multiset [of M] | |
| 174 | by (simp add: multiset_def set_of_def count_def) | |
| 10249 | 175 | |
| 17161 | 176 | lemma setsum_count_Int: | 
| 11464 | 177 | "finite A ==> setsum (count N) (A \<inter> set_of N) = setsum (count N) A" | 
| 18258 | 178 | apply (induct rule: finite_induct) | 
| 17161 | 179 | apply simp | 
| 10249 | 180 | apply (simp add: Int_insert_left set_of_def) | 
| 181 | done | |
| 182 | ||
| 17161 | 183 | lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" | 
| 10249 | 184 | apply (unfold size_def) | 
| 11464 | 185 | apply (subgoal_tac "count (M + N) = (\<lambda>a. count M a + count N a)") | 
| 10249 | 186 | prefer 2 | 
| 15072 | 187 | apply (rule ext, simp) | 
| 15402 | 188 | apply (simp (no_asm_simp) add: setsum_Un_nat setsum_addf setsum_count_Int) | 
| 10249 | 189 | apply (subst Int_commute) | 
| 190 | apply (simp (no_asm_simp) add: setsum_count_Int) | |
| 191 | done | |
| 192 | ||
| 17161 | 193 | lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
 | 
| 15072 | 194 | apply (unfold size_def Mempty_def count_def, auto) | 
| 10249 | 195 | apply (simp add: set_of_def count_def expand_fun_eq) | 
| 196 | done | |
| 197 | ||
| 17161 | 198 | lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \<exists>a. a :# M" | 
| 10249 | 199 | apply (unfold size_def) | 
| 15072 | 200 | apply (drule setsum_SucD, auto) | 
| 10249 | 201 | done | 
| 202 | ||
| 203 | ||
| 204 | subsubsection {* Equality of multisets *}
 | |
| 205 | ||
| 17161 | 206 | lemma multiset_eq_conv_count_eq: "(M = N) = (\<forall>a. count M a = count N a)" | 
| 207 | by (simp add: count_def expand_fun_eq) | |
| 10249 | 208 | |
| 17161 | 209 | lemma single_not_empty [simp]: "{#a#} \<noteq> {#} \<and> {#} \<noteq> {#a#}"
 | 
| 210 | by (simp add: single_def Mempty_def expand_fun_eq) | |
| 10249 | 211 | |
| 17161 | 212 | lemma single_eq_single [simp]: "({#a#} = {#b#}) = (a = b)"
 | 
| 213 | by (auto simp add: single_def expand_fun_eq) | |
| 10249 | 214 | |
| 17161 | 215 | lemma union_eq_empty [iff]: "(M + N = {#}) = (M = {#} \<and> N = {#})"
 | 
| 216 | by (auto simp add: union_def Mempty_def expand_fun_eq) | |
| 10249 | 217 | |
| 17161 | 218 | lemma empty_eq_union [iff]: "({#} = M + N) = (M = {#} \<and> N = {#})"
 | 
| 219 | by (auto simp add: union_def Mempty_def expand_fun_eq) | |
| 10249 | 220 | |
| 17161 | 221 | lemma union_right_cancel [simp]: "(M + K = N + K) = (M = (N::'a multiset))" | 
| 222 | by (simp add: union_def expand_fun_eq) | |
| 10249 | 223 | |
| 17161 | 224 | lemma union_left_cancel [simp]: "(K + M = K + N) = (M = (N::'a multiset))" | 
| 225 | by (simp add: union_def expand_fun_eq) | |
| 10249 | 226 | |
| 17161 | 227 | lemma union_is_single: | 
| 11464 | 228 |     "(M + N = {#a#}) = (M = {#a#} \<and> N={#} \<or> M = {#} \<and> N = {#a#})"
 | 
| 15072 | 229 | apply (simp add: Mempty_def single_def union_def add_is_1 expand_fun_eq) | 
| 10249 | 230 | apply blast | 
| 231 | done | |
| 232 | ||
| 17161 | 233 | lemma single_is_union: | 
| 15072 | 234 |      "({#a#} = M + N) = ({#a#} = M \<and> N = {#} \<or> M = {#} \<and> {#a#} = N)"
 | 
| 10249 | 235 | apply (unfold Mempty_def single_def union_def) | 
| 11464 | 236 | apply (simp add: add_is_1 one_is_add expand_fun_eq) | 
| 10249 | 237 | apply (blast dest: sym) | 
| 238 | done | |
| 239 | ||
| 17778 | 240 | ML"reset use_neq_simproc" | 
| 17161 | 241 | lemma add_eq_conv_diff: | 
| 10249 | 242 |   "(M + {#a#} = N + {#b#}) =
 | 
| 15072 | 243 |    (M = N \<and> a = b \<or> M = N - {#a#} + {#b#} \<and> N = M - {#b#} + {#a#})"
 | 
| 10249 | 244 | apply (unfold single_def union_def diff_def) | 
| 245 | apply (simp (no_asm) add: expand_fun_eq) | |
| 15072 | 246 | apply (rule conjI, force, safe, simp_all) | 
| 13601 | 247 | apply (simp add: eq_sym_conv) | 
| 10249 | 248 | done | 
| 17778 | 249 | ML"set use_neq_simproc" | 
| 10249 | 250 | |
| 15869 | 251 | declare Rep_multiset_inject [symmetric, simp del] | 
| 252 | ||
| 253 | ||
| 254 | subsubsection {* Intersection *}
 | |
| 255 | ||
| 256 | lemma multiset_inter_count: | |
| 17161 | 257 | "count (A #\<inter> B) x = min (count A x) (count B x)" | 
| 258 | by (simp add: multiset_inter_def min_def) | |
| 15869 | 259 | |
| 260 | lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A" | |
| 17200 | 261 | by (simp add: multiset_eq_conv_count_eq multiset_inter_count | 
| 17161 | 262 | min_max.below_inf.inf_commute) | 
| 15869 | 263 | |
| 264 | lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C" | |
| 17200 | 265 | by (simp add: multiset_eq_conv_count_eq multiset_inter_count | 
| 17161 | 266 | min_max.below_inf.inf_assoc) | 
| 15869 | 267 | |
| 268 | lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)" | |
| 269 | by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def) | |
| 270 | ||
| 17161 | 271 | lemmas multiset_inter_ac = | 
| 272 | multiset_inter_commute | |
| 273 | multiset_inter_assoc | |
| 274 | multiset_inter_left_commute | |
| 15869 | 275 | |
| 276 | lemma multiset_union_diff_commute: "B #\<inter> C = {#} \<Longrightarrow> A + B - C = A - C + B"
 | |
| 17200 | 277 | apply (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def | 
| 17161 | 278 | split: split_if_asm) | 
| 15869 | 279 | apply clarsimp | 
| 17161 | 280 | apply (erule_tac x = a in allE) | 
| 15869 | 281 | apply auto | 
| 282 | done | |
| 283 | ||
| 10249 | 284 | |
| 285 | subsection {* Induction over multisets *}
 | |
| 286 | ||
| 287 | lemma setsum_decr: | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11655diff
changeset | 288 | "finite F ==> (0::nat) < f a ==> | 
| 15072 | 289 | setsum (f (a := f a - 1)) F = (if a\<in>F then setsum f F - 1 else setsum f F)" | 
| 18258 | 290 | apply (induct rule: finite_induct) | 
| 291 | apply auto | |
| 15072 | 292 | apply (drule_tac a = a in mk_disjoint_insert, auto) | 
| 10249 | 293 | done | 
| 294 | ||
| 10313 | 295 | lemma rep_multiset_induct_aux: | 
| 18730 | 296 | assumes 1: "P (\<lambda>a. (0::nat))" | 
| 297 | and 2: "!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))" | |
| 17161 | 298 |   shows "\<forall>f. f \<in> multiset --> setsum f {x. 0 < f x} = n --> P f"
 | 
| 18730 | 299 | apply (unfold multiset_def) | 
| 300 | apply (induct_tac n, simp, clarify) | |
| 301 | apply (subgoal_tac "f = (\<lambda>a.0)") | |
| 302 | apply simp | |
| 303 | apply (rule 1) | |
| 304 | apply (rule ext, force, clarify) | |
| 305 | apply (frule setsum_SucD, clarify) | |
| 306 | apply (rename_tac a) | |
| 307 |   apply (subgoal_tac "finite {x. 0 < (f (a := f a - 1)) x}")
 | |
| 308 | prefer 2 | |
| 309 | apply (rule finite_subset) | |
| 310 | prefer 2 | |
| 311 | apply assumption | |
| 312 | apply simp | |
| 313 | apply blast | |
| 314 | apply (subgoal_tac "f = (f (a := f a - 1))(a := (f (a := f a - 1)) a + 1)") | |
| 315 | prefer 2 | |
| 316 | apply (rule ext) | |
| 317 | apply (simp (no_asm_simp)) | |
| 318 | apply (erule ssubst, rule 2 [unfolded multiset_def], blast) | |
| 319 | apply (erule allE, erule impE, erule_tac [2] mp, blast) | |
| 320 | apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply One_nat_def) | |
| 321 |   apply (subgoal_tac "{x. x \<noteq> a --> 0 < f x} = {x. 0 < f x}")
 | |
| 322 | prefer 2 | |
| 323 | apply blast | |
| 324 |   apply (subgoal_tac "{x. x \<noteq> a \<and> 0 < f x} = {x. 0 < f x} - {a}")
 | |
| 325 | prefer 2 | |
| 326 | apply blast | |
| 327 | apply (simp add: le_imp_diff_is_add setsum_diff1_nat cong: conj_cong) | |
| 328 | done | |
| 10249 | 329 | |
| 10313 | 330 | theorem rep_multiset_induct: | 
| 11464 | 331 | "f \<in> multiset ==> P (\<lambda>a. 0) ==> | 
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11655diff
changeset | 332 | (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))) ==> P f" | 
| 17161 | 333 | using rep_multiset_induct_aux by blast | 
| 10249 | 334 | |
| 18258 | 335 | theorem multiset_induct [case_names empty add, induct type: multiset]: | 
| 336 |   assumes empty: "P {#}"
 | |
| 337 |     and add: "!!M x. P M ==> P (M + {#x#})"
 | |
| 17161 | 338 | shows "P M" | 
| 10249 | 339 | proof - | 
| 340 | note defns = union_def single_def Mempty_def | |
| 341 | show ?thesis | |
| 342 | apply (rule Rep_multiset_inverse [THEN subst]) | |
| 10313 | 343 | apply (rule Rep_multiset [THEN rep_multiset_induct]) | 
| 18258 | 344 | apply (rule empty [unfolded defns]) | 
| 15072 | 345 | apply (subgoal_tac "f(b := f b + 1) = (\<lambda>a. f a + (if a=b then 1 else 0))") | 
| 10249 | 346 | prefer 2 | 
| 347 | apply (simp add: expand_fun_eq) | |
| 348 | apply (erule ssubst) | |
| 17200 | 349 | apply (erule Abs_multiset_inverse [THEN subst]) | 
| 18258 | 350 | apply (erule add [unfolded defns, simplified]) | 
| 10249 | 351 | done | 
| 352 | qed | |
| 353 | ||
| 354 | lemma MCollect_preserves_multiset: | |
| 11464 | 355 | "M \<in> multiset ==> (\<lambda>x. if P x then M x else 0) \<in> multiset" | 
| 10249 | 356 | apply (simp add: multiset_def) | 
| 15072 | 357 | apply (rule finite_subset, auto) | 
| 10249 | 358 | done | 
| 359 | ||
| 17161 | 360 | lemma count_MCollect [simp]: | 
| 10249 | 361 |     "count {# x:M. P x #} a = (if P a then count M a else 0)"
 | 
| 15072 | 362 | by (simp add: count_def MCollect_def MCollect_preserves_multiset) | 
| 10249 | 363 | |
| 17161 | 364 | lemma set_of_MCollect [simp]: "set_of {# x:M. P x #} = set_of M \<inter> {x. P x}"
 | 
| 365 | by (auto simp add: set_of_def) | |
| 10249 | 366 | |
| 17161 | 367 | lemma multiset_partition: "M = {# x:M. P x #} + {# x:M. \<not> P x #}"
 | 
| 368 | by (subst multiset_eq_conv_count_eq, auto) | |
| 10249 | 369 | |
| 17161 | 370 | lemma add_eq_conv_ex: | 
| 371 |   "(M + {#a#} = N + {#b#}) =
 | |
| 372 |     (M = N \<and> a = b \<or> (\<exists>K. M = K + {#b#} \<and> N = K + {#a#}))"
 | |
| 15072 | 373 | by (auto simp add: add_eq_conv_diff) | 
| 10249 | 374 | |
| 15869 | 375 | declare multiset_typedef [simp del] | 
| 10249 | 376 | |
| 17161 | 377 | |
| 10249 | 378 | subsection {* Multiset orderings *}
 | 
| 379 | ||
| 380 | subsubsection {* Well-foundedness *}
 | |
| 381 | ||
| 19086 | 382 | definition | 
| 11464 | 383 |   mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set"
 | 
| 19086 | 384 | "mult1 r = | 
| 11464 | 385 |     {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
 | 
| 386 | (\<forall>b. b :# K --> (b, a) \<in> r)}" | |
| 10249 | 387 | |
| 11464 | 388 |   mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set"
 | 
| 19086 | 389 | "mult r = (mult1 r)\<^sup>+" | 
| 10249 | 390 | |
| 11464 | 391 | lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
 | 
| 10277 | 392 | by (simp add: mult1_def) | 
| 10249 | 393 | |
| 11464 | 394 | lemma less_add: "(N, M0 + {#a#}) \<in> mult1 r ==>
 | 
| 395 |     (\<exists>M. (M, M0) \<in> mult1 r \<and> N = M + {#a#}) \<or>
 | |
| 396 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K)" | |
| 19582 | 397 | (is "_ \<Longrightarrow> ?case1 (mult1 r) \<or> ?case2") | 
| 10249 | 398 | proof (unfold mult1_def) | 
| 11464 | 399 | let ?r = "\<lambda>K a. \<forall>b. b :# K --> (b, a) \<in> r" | 
| 400 |   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
 | |
| 10249 | 401 |   let ?case1 = "?case1 {(N, M). ?R N M}"
 | 
| 402 | ||
| 11464 | 403 |   assume "(N, M0 + {#a#}) \<in> {(N, M). ?R N M}"
 | 
| 18258 | 404 | then have "\<exists>a' M0' K. | 
| 11464 | 405 |       M0 + {#a#} = M0' + {#a'#} \<and> N = M0' + K \<and> ?r K a'" by simp
 | 
| 18258 | 406 | then show "?case1 \<or> ?case2" | 
| 10249 | 407 | proof (elim exE conjE) | 
| 408 | fix a' M0' K | |
| 409 | assume N: "N = M0' + K" and r: "?r K a'" | |
| 410 |     assume "M0 + {#a#} = M0' + {#a'#}"
 | |
| 18258 | 411 | then have "M0 = M0' \<and> a = a' \<or> | 
| 11464 | 412 |         (\<exists>K'. M0 = K' + {#a'#} \<and> M0' = K' + {#a#})"
 | 
| 10249 | 413 | by (simp only: add_eq_conv_ex) | 
| 18258 | 414 | then show ?thesis | 
| 10249 | 415 | proof (elim disjE conjE exE) | 
| 416 | assume "M0 = M0'" "a = a'" | |
| 11464 | 417 | with N r have "?r K a \<and> N = M0 + K" by simp | 
| 18258 | 418 | then have ?case2 .. then show ?thesis .. | 
| 10249 | 419 | next | 
| 420 | fix K' | |
| 421 |       assume "M0' = K' + {#a#}"
 | |
| 422 |       with N have n: "N = K' + K + {#a#}" by (simp add: union_ac)
 | |
| 423 | ||
| 424 |       assume "M0 = K' + {#a'#}"
 | |
| 425 | with r have "?R (K' + K) M0" by blast | |
| 18258 | 426 | with n have ?case1 by simp then show ?thesis .. | 
| 10249 | 427 | qed | 
| 428 | qed | |
| 429 | qed | |
| 430 | ||
| 11464 | 431 | lemma all_accessible: "wf r ==> \<forall>M. M \<in> acc (mult1 r)" | 
| 10249 | 432 | proof | 
| 433 | let ?R = "mult1 r" | |
| 434 | let ?W = "acc ?R" | |
| 435 |   {
 | |
| 436 | fix M M0 a | |
| 11464 | 437 | assume M0: "M0 \<in> ?W" | 
| 12399 | 438 |       and wf_hyp: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | 
| 11464 | 439 |       and acc_hyp: "\<forall>M. (M, M0) \<in> ?R --> M + {#a#} \<in> ?W"
 | 
| 440 |     have "M0 + {#a#} \<in> ?W"
 | |
| 10249 | 441 |     proof (rule accI [of "M0 + {#a#}"])
 | 
| 442 | fix N | |
| 11464 | 443 |       assume "(N, M0 + {#a#}) \<in> ?R"
 | 
| 18258 | 444 |       then have "((\<exists>M. (M, M0) \<in> ?R \<and> N = M + {#a#}) \<or>
 | 
| 11464 | 445 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K))" | 
| 10249 | 446 | by (rule less_add) | 
| 18258 | 447 | then show "N \<in> ?W" | 
| 10249 | 448 | proof (elim exE disjE conjE) | 
| 11464 | 449 |         fix M assume "(M, M0) \<in> ?R" and N: "N = M + {#a#}"
 | 
| 450 |         from acc_hyp have "(M, M0) \<in> ?R --> M + {#a#} \<in> ?W" ..
 | |
| 18258 | 451 |         then have "M + {#a#} \<in> ?W" ..
 | 
| 452 | then show "N \<in> ?W" by (simp only: N) | |
| 10249 | 453 | next | 
| 454 | fix K | |
| 455 | assume N: "N = M0 + K" | |
| 11464 | 456 | assume "\<forall>b. b :# K --> (b, a) \<in> r" | 
| 18730 | 457 | then have "M0 + K \<in> ?W" | 
| 10249 | 458 | proof (induct K) | 
| 18730 | 459 | case empty | 
| 18258 | 460 |           from M0 show "M0 + {#} \<in> ?W" by simp
 | 
| 18730 | 461 | next | 
| 462 | case (add K x) | |
| 463 | from add.prems have "(x, a) \<in> r" by simp | |
| 18258 | 464 |           with wf_hyp have "\<forall>M \<in> ?W. M + {#x#} \<in> ?W" by blast
 | 
| 18730 | 465 | moreover from add have "M0 + K \<in> ?W" by simp | 
| 18258 | 466 |           ultimately have "(M0 + K) + {#x#} \<in> ?W" ..
 | 
| 467 |           then show "M0 + (K + {#x#}) \<in> ?W" by (simp only: union_assoc)
 | |
| 10249 | 468 | qed | 
| 18730 | 469 | then show "N \<in> ?W" by (simp only: N) | 
| 10249 | 470 | qed | 
| 471 | qed | |
| 472 | } note tedious_reasoning = this | |
| 473 | ||
| 474 | assume wf: "wf r" | |
| 475 | fix M | |
| 11464 | 476 | show "M \<in> ?W" | 
| 10249 | 477 | proof (induct M) | 
| 11464 | 478 |     show "{#} \<in> ?W"
 | 
| 10249 | 479 | proof (rule accI) | 
| 11464 | 480 |       fix b assume "(b, {#}) \<in> ?R"
 | 
| 481 | with not_less_empty show "b \<in> ?W" by contradiction | |
| 10249 | 482 | qed | 
| 483 | ||
| 11464 | 484 | fix M a assume "M \<in> ?W" | 
| 485 |     from wf have "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | |
| 10249 | 486 | proof induct | 
| 487 | fix a | |
| 12399 | 488 |       assume "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | 
| 11464 | 489 |       show "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | 
| 10249 | 490 | proof | 
| 11464 | 491 | fix M assume "M \<in> ?W" | 
| 18258 | 492 |         then show "M + {#a#} \<in> ?W"
 | 
| 10249 | 493 | by (rule acc_induct) (rule tedious_reasoning) | 
| 494 | qed | |
| 495 | qed | |
| 18258 | 496 |     then show "M + {#a#} \<in> ?W" ..
 | 
| 10249 | 497 | qed | 
| 498 | qed | |
| 499 | ||
| 500 | theorem wf_mult1: "wf r ==> wf (mult1 r)" | |
| 501 | by (rule acc_wfI, rule all_accessible) | |
| 502 | ||
| 503 | theorem wf_mult: "wf r ==> wf (mult r)" | |
| 504 | by (unfold mult_def, rule wf_trancl, rule wf_mult1) | |
| 505 | ||
| 506 | ||
| 507 | subsubsection {* Closure-free presentation *}
 | |
| 508 | ||
| 509 | (*Badly needed: a linear arithmetic procedure for multisets*) | |
| 510 | ||
| 511 | lemma diff_union_single_conv: "a :# J ==> I + J - {#a#} = I + (J - {#a#})"
 | |
| 15072 | 512 | by (simp add: multiset_eq_conv_count_eq) | 
| 10249 | 513 | |
| 514 | text {* One direction. *}
 | |
| 515 | ||
| 516 | lemma mult_implies_one_step: | |
| 11464 | 517 | "trans r ==> (M, N) \<in> mult r ==> | 
| 518 |     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
 | |
| 519 | (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r)" | |
| 10249 | 520 | apply (unfold mult_def mult1_def set_of_def) | 
| 15072 | 521 | apply (erule converse_trancl_induct, clarify) | 
| 522 | apply (rule_tac x = M0 in exI, simp, clarify) | |
| 10249 | 523 | apply (case_tac "a :# K") | 
| 524 | apply (rule_tac x = I in exI) | |
| 525 | apply (simp (no_asm)) | |
| 526 |    apply (rule_tac x = "(K - {#a#}) + Ka" in exI)
 | |
| 527 | apply (simp (no_asm_simp) add: union_assoc [symmetric]) | |
| 11464 | 528 |    apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong)
 | 
| 10249 | 529 | apply (simp add: diff_union_single_conv) | 
| 530 | apply (simp (no_asm_use) add: trans_def) | |
| 531 | apply blast | |
| 532 | apply (subgoal_tac "a :# I") | |
| 533 |    apply (rule_tac x = "I - {#a#}" in exI)
 | |
| 534 |    apply (rule_tac x = "J + {#a#}" in exI)
 | |
| 535 | apply (rule_tac x = "K + Ka" in exI) | |
| 536 | apply (rule conjI) | |
| 537 | apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) | |
| 538 | apply (rule conjI) | |
| 15072 | 539 |     apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong, simp)
 | 
| 10249 | 540 | apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) | 
| 541 | apply (simp (no_asm_use) add: trans_def) | |
| 542 | apply blast | |
| 10277 | 543 |   apply (subgoal_tac "a :# (M0 + {#a#})")
 | 
| 10249 | 544 | apply simp | 
| 545 | apply (simp (no_asm)) | |
| 546 | done | |
| 547 | ||
| 548 | lemma elem_imp_eq_diff_union: "a :# M ==> M = M - {#a#} + {#a#}"
 | |
| 15072 | 549 | by (simp add: multiset_eq_conv_count_eq) | 
| 10249 | 550 | |
| 11464 | 551 | lemma size_eq_Suc_imp_eq_union: "size M = Suc n ==> \<exists>a N. M = N + {#a#}"
 | 
| 10249 | 552 | apply (erule size_eq_Suc_imp_elem [THEN exE]) | 
| 15072 | 553 | apply (drule elem_imp_eq_diff_union, auto) | 
| 10249 | 554 | done | 
| 555 | ||
| 556 | lemma one_step_implies_mult_aux: | |
| 557 | "trans r ==> | |
| 11464 | 558 |     \<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))
 | 
| 559 | --> (I + K, I + J) \<in> mult r" | |
| 15072 | 560 | apply (induct_tac n, auto) | 
| 561 | apply (frule size_eq_Suc_imp_eq_union, clarify) | |
| 562 | apply (rename_tac "J'", simp) | |
| 563 | apply (erule notE, auto) | |
| 10249 | 564 |   apply (case_tac "J' = {#}")
 | 
| 565 | apply (simp add: mult_def) | |
| 566 | apply (rule r_into_trancl) | |
| 15072 | 567 | apply (simp add: mult1_def set_of_def, blast) | 
| 11464 | 568 |   txt {* Now we know @{term "J' \<noteq> {#}"}. *}
 | 
| 569 | apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition) | |
| 570 | apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp) | |
| 10249 | 571 | apply (erule ssubst) | 
| 15072 | 572 | apply (simp add: Ball_def, auto) | 
| 10249 | 573 | apply (subgoal_tac | 
| 11464 | 574 |     "((I + {# x : K. (x, a) \<in> r #}) + {# x : K. (x, a) \<notin> r #},
 | 
| 575 |       (I + {# x : K. (x, a) \<in> r #}) + J') \<in> mult r")
 | |
| 10249 | 576 | prefer 2 | 
| 577 | apply force | |
| 578 | apply (simp (no_asm_use) add: union_assoc [symmetric] mult_def) | |
| 579 | apply (erule trancl_trans) | |
| 580 | apply (rule r_into_trancl) | |
| 581 | apply (simp add: mult1_def set_of_def) | |
| 582 | apply (rule_tac x = a in exI) | |
| 583 | apply (rule_tac x = "I + J'" in exI) | |
| 584 | apply (simp add: union_ac) | |
| 585 | done | |
| 586 | ||
| 17161 | 587 | lemma one_step_implies_mult: | 
| 11464 | 588 |   "trans r ==> J \<noteq> {#} ==> \<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r
 | 
| 589 | ==> (I + K, I + J) \<in> mult r" | |
| 15072 | 590 | apply (insert one_step_implies_mult_aux, blast) | 
| 10249 | 591 | done | 
| 592 | ||
| 593 | ||
| 594 | subsubsection {* Partial-order properties *}
 | |
| 595 | ||
| 12338 
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
 wenzelm parents: 
11868diff
changeset | 596 | instance multiset :: (type) ord .. | 
| 10249 | 597 | |
| 598 | defs (overloaded) | |
| 11464 | 599 |   less_multiset_def: "M' < M == (M', M) \<in> mult {(x', x). x' < x}"
 | 
| 600 | le_multiset_def: "M' <= M == M' = M \<or> M' < (M::'a multiset)" | |
| 10249 | 601 | |
| 602 | lemma trans_base_order: "trans {(x', x). x' < (x::'a::order)}"
 | |
| 18730 | 603 | unfolding trans_def by (blast intro: order_less_trans) | 
| 10249 | 604 | |
| 605 | text {*
 | |
| 606 | \medskip Irreflexivity. | |
| 607 | *} | |
| 608 | ||
| 609 | lemma mult_irrefl_aux: | |
| 18258 | 610 |     "finite A ==> (\<forall>x \<in> A. \<exists>y \<in> A. x < (y::'a::order)) \<Longrightarrow> A = {}"
 | 
| 611 | apply (induct rule: finite_induct) | |
| 10249 | 612 | apply (auto intro: order_less_trans) | 
| 613 | done | |
| 614 | ||
| 17161 | 615 | lemma mult_less_not_refl: "\<not> M < (M::'a::order multiset)" | 
| 15072 | 616 | apply (unfold less_multiset_def, auto) | 
| 617 | apply (drule trans_base_order [THEN mult_implies_one_step], auto) | |
| 10249 | 618 | apply (drule finite_set_of [THEN mult_irrefl_aux [rule_format (no_asm)]]) | 
| 619 | apply (simp add: set_of_eq_empty_iff) | |
| 620 | done | |
| 621 | ||
| 622 | lemma mult_less_irrefl [elim!]: "M < (M::'a::order multiset) ==> R" | |
| 15072 | 623 | by (insert mult_less_not_refl, fast) | 
| 10249 | 624 | |
| 625 | ||
| 626 | text {* Transitivity. *}
 | |
| 627 | ||
| 628 | theorem mult_less_trans: "K < M ==> M < N ==> K < (N::'a::order multiset)" | |
| 629 | apply (unfold less_multiset_def mult_def) | |
| 630 | apply (blast intro: trancl_trans) | |
| 631 | done | |
| 632 | ||
| 633 | text {* Asymmetry. *}
 | |
| 634 | ||
| 11464 | 635 | theorem mult_less_not_sym: "M < N ==> \<not> N < (M::'a::order multiset)" | 
| 10249 | 636 | apply auto | 
| 637 | apply (rule mult_less_not_refl [THEN notE]) | |
| 15072 | 638 | apply (erule mult_less_trans, assumption) | 
| 10249 | 639 | done | 
| 640 | ||
| 641 | theorem mult_less_asym: | |
| 11464 | 642 | "M < N ==> (\<not> P ==> N < (M::'a::order multiset)) ==> P" | 
| 15072 | 643 | by (insert mult_less_not_sym, blast) | 
| 10249 | 644 | |
| 645 | theorem mult_le_refl [iff]: "M <= (M::'a::order multiset)" | |
| 18730 | 646 | unfolding le_multiset_def by auto | 
| 10249 | 647 | |
| 648 | text {* Anti-symmetry. *}
 | |
| 649 | ||
| 650 | theorem mult_le_antisym: | |
| 651 | "M <= N ==> N <= M ==> M = (N::'a::order multiset)" | |
| 18730 | 652 | unfolding le_multiset_def by (blast dest: mult_less_not_sym) | 
| 10249 | 653 | |
| 654 | text {* Transitivity. *}
 | |
| 655 | ||
| 656 | theorem mult_le_trans: | |
| 657 | "K <= M ==> M <= N ==> K <= (N::'a::order multiset)" | |
| 18730 | 658 | unfolding le_multiset_def by (blast intro: mult_less_trans) | 
| 10249 | 659 | |
| 11655 | 660 | theorem mult_less_le: "(M < N) = (M <= N \<and> M \<noteq> (N::'a::order multiset))" | 
| 18730 | 661 | unfolding le_multiset_def by auto | 
| 10249 | 662 | |
| 10277 | 663 | text {* Partial order. *}
 | 
| 664 | ||
| 665 | instance multiset :: (order) order | |
| 666 | apply intro_classes | |
| 667 | apply (rule mult_le_refl) | |
| 15072 | 668 | apply (erule mult_le_trans, assumption) | 
| 669 | apply (erule mult_le_antisym, assumption) | |
| 10277 | 670 | apply (rule mult_less_le) | 
| 671 | done | |
| 672 | ||
| 10249 | 673 | |
| 674 | subsubsection {* Monotonicity of multiset union *}
 | |
| 675 | ||
| 17161 | 676 | lemma mult1_union: | 
| 11464 | 677 | "(B, D) \<in> mult1 r ==> trans r ==> (C + B, C + D) \<in> mult1 r" | 
| 15072 | 678 | apply (unfold mult1_def, auto) | 
| 10249 | 679 | apply (rule_tac x = a in exI) | 
| 680 | apply (rule_tac x = "C + M0" in exI) | |
| 681 | apply (simp add: union_assoc) | |
| 682 | done | |
| 683 | ||
| 684 | lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)" | |
| 685 | apply (unfold less_multiset_def mult_def) | |
| 686 | apply (erule trancl_induct) | |
| 687 | apply (blast intro: mult1_union transI order_less_trans r_into_trancl) | |
| 688 | apply (blast intro: mult1_union transI order_less_trans r_into_trancl trancl_trans) | |
| 689 | done | |
| 690 | ||
| 691 | lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)" | |
| 692 | apply (subst union_commute [of B C]) | |
| 693 | apply (subst union_commute [of D C]) | |
| 694 | apply (erule union_less_mono2) | |
| 695 | done | |
| 696 | ||
| 17161 | 697 | lemma union_less_mono: | 
| 10249 | 698 | "A < C ==> B < D ==> A + B < C + (D::'a::order multiset)" | 
| 699 | apply (blast intro!: union_less_mono1 union_less_mono2 mult_less_trans) | |
| 700 | done | |
| 701 | ||
| 17161 | 702 | lemma union_le_mono: | 
| 10249 | 703 | "A <= C ==> B <= D ==> A + B <= C + (D::'a::order multiset)" | 
| 18730 | 704 | unfolding le_multiset_def | 
| 705 | by (blast intro: union_less_mono union_less_mono1 union_less_mono2) | |
| 10249 | 706 | |
| 17161 | 707 | lemma empty_leI [iff]: "{#} <= (M::'a::order multiset)"
 | 
| 10249 | 708 | apply (unfold le_multiset_def less_multiset_def) | 
| 709 |   apply (case_tac "M = {#}")
 | |
| 710 | prefer 2 | |
| 11464 | 711 |    apply (subgoal_tac "({#} + {#}, {#} + M) \<in> mult (Collect (split op <))")
 | 
| 10249 | 712 | prefer 2 | 
| 713 | apply (rule one_step_implies_mult) | |
| 15072 | 714 | apply (simp only: trans_def, auto) | 
| 10249 | 715 | done | 
| 716 | ||
| 17161 | 717 | lemma union_upper1: "A <= A + (B::'a::order multiset)" | 
| 15072 | 718 | proof - | 
| 17200 | 719 |   have "A + {#} <= A + B" by (blast intro: union_le_mono)
 | 
| 18258 | 720 | then show ?thesis by simp | 
| 15072 | 721 | qed | 
| 722 | ||
| 17161 | 723 | lemma union_upper2: "B <= A + (B::'a::order multiset)" | 
| 18258 | 724 | by (subst union_commute) (rule union_upper1) | 
| 15072 | 725 | |
| 726 | ||
| 17200 | 727 | subsection {* Link with lists *}
 | 
| 15072 | 728 | |
| 17200 | 729 | consts | 
| 15072 | 730 | multiset_of :: "'a list \<Rightarrow> 'a multiset" | 
| 731 | primrec | |
| 732 |   "multiset_of [] = {#}"
 | |
| 733 |   "multiset_of (a # x) = multiset_of x + {# a #}"
 | |
| 734 | ||
| 735 | lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])"
 | |
| 18258 | 736 | by (induct x) auto | 
| 15072 | 737 | |
| 738 | lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])"
 | |
| 18258 | 739 | by (induct x) auto | 
| 15072 | 740 | |
| 741 | lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" | |
| 18258 | 742 | by (induct x) auto | 
| 15867 | 743 | |
| 744 | lemma mem_set_multiset_eq: "x \<in> set xs = (x :# multiset_of xs)" | |
| 745 | by (induct xs) auto | |
| 15072 | 746 | |
| 18258 | 747 | lemma multiset_of_append [simp]: | 
| 748 | "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" | |
| 749 | by (induct xs fixing: ys) (auto simp: union_ac) | |
| 18730 | 750 | |
| 15072 | 751 | lemma surj_multiset_of: "surj multiset_of" | 
| 17200 | 752 | apply (unfold surj_def, rule allI) | 
| 753 | apply (rule_tac M=y in multiset_induct, auto) | |
| 754 | apply (rule_tac x = "x # xa" in exI, auto) | |
| 10249 | 755 | done | 
| 756 | ||
| 15072 | 757 | lemma set_count_greater_0: "set x = {a. 0 < count (multiset_of x) a}"
 | 
| 18258 | 758 | by (induct x) auto | 
| 15072 | 759 | |
| 17200 | 760 | lemma distinct_count_atmost_1: | 
| 15072 | 761 | "distinct x = (! a. count (multiset_of x) a = (if a \<in> set x then 1 else 0))" | 
| 18258 | 762 | apply (induct x, simp, rule iffI, simp_all) | 
| 17200 | 763 | apply (rule conjI) | 
| 764 | apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) | |
| 15072 | 765 | apply (erule_tac x=a in allE, simp, clarify) | 
| 17200 | 766 | apply (erule_tac x=aa in allE, simp) | 
| 15072 | 767 | done | 
| 768 | ||
| 17200 | 769 | lemma multiset_of_eq_setD: | 
| 15867 | 770 | "multiset_of xs = multiset_of ys \<Longrightarrow> set xs = set ys" | 
| 771 | by (rule) (auto simp add:multiset_eq_conv_count_eq set_count_greater_0) | |
| 772 | ||
| 17200 | 773 | lemma set_eq_iff_multiset_of_eq_distinct: | 
| 774 | "\<lbrakk>distinct x; distinct y\<rbrakk> | |
| 15072 | 775 | \<Longrightarrow> (set x = set y) = (multiset_of x = multiset_of y)" | 
| 17200 | 776 | by (auto simp: multiset_eq_conv_count_eq distinct_count_atmost_1) | 
| 15072 | 777 | |
| 17200 | 778 | lemma set_eq_iff_multiset_of_remdups_eq: | 
| 15072 | 779 | "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" | 
| 17200 | 780 | apply (rule iffI) | 
| 781 | apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) | |
| 782 | apply (drule distinct_remdups[THEN distinct_remdups | |
| 783 | [THEN set_eq_iff_multiset_of_eq_distinct[THEN iffD2]]]) | |
| 15072 | 784 | apply simp | 
| 10249 | 785 | done | 
| 786 | ||
| 18258 | 787 | lemma multiset_of_compl_union [simp]: | 
| 788 | "multiset_of [x\<in>xs. P x] + multiset_of [x\<in>xs. \<not>P x] = multiset_of xs" | |
| 15630 | 789 | by (induct xs) (auto simp: union_ac) | 
| 15072 | 790 | |
| 17200 | 791 | lemma count_filter: | 
| 18258 | 792 | "count (multiset_of xs) x = length [y \<in> xs. y = x]" | 
| 793 | by (induct xs) auto | |
| 15867 | 794 | |
| 795 | ||
| 15072 | 796 | subsection {* Pointwise ordering induced by count *}
 | 
| 797 | ||
| 19086 | 798 | definition | 
| 799 |   mset_le :: "['a multiset, 'a multiset] \<Rightarrow> bool"   ("_ \<le># _"  [50,51] 50)
 | |
| 800 | "(xs \<le># ys) = (\<forall>a. count xs a \<le> count ys a)" | |
| 15072 | 801 | |
| 802 | lemma mset_le_refl[simp]: "xs \<le># xs" | |
| 18730 | 803 | unfolding mset_le_def by auto | 
| 15072 | 804 | |
| 805 | lemma mset_le_trans: "\<lbrakk> xs \<le># ys; ys \<le># zs \<rbrakk> \<Longrightarrow> xs \<le># zs" | |
| 18730 | 806 | unfolding mset_le_def by (fast intro: order_trans) | 
| 15072 | 807 | |
| 808 | lemma mset_le_antisym: "\<lbrakk> xs\<le># ys; ys \<le># xs\<rbrakk> \<Longrightarrow> xs = ys" | |
| 17200 | 809 | apply (unfold mset_le_def) | 
| 810 | apply (rule multiset_eq_conv_count_eq[THEN iffD2]) | |
| 15072 | 811 | apply (blast intro: order_antisym) | 
| 812 | done | |
| 813 | ||
| 17200 | 814 | lemma mset_le_exists_conv: | 
| 815 | "(xs \<le># ys) = (\<exists>zs. ys = xs + zs)" | |
| 816 | apply (unfold mset_le_def, rule iffI, rule_tac x = "ys - xs" in exI) | |
| 15072 | 817 | apply (auto intro: multiset_eq_conv_count_eq [THEN iffD2]) | 
| 818 | done | |
| 819 | ||
| 820 | lemma mset_le_mono_add_right_cancel[simp]: "(xs + zs \<le># ys + zs) = (xs \<le># ys)" | |
| 18730 | 821 | unfolding mset_le_def by auto | 
| 15072 | 822 | |
| 823 | lemma mset_le_mono_add_left_cancel[simp]: "(zs + xs \<le># zs + ys) = (xs \<le># ys)" | |
| 18730 | 824 | unfolding mset_le_def by auto | 
| 15072 | 825 | |
| 17200 | 826 | lemma mset_le_mono_add: "\<lbrakk> xs \<le># ys; vs \<le># ws \<rbrakk> \<Longrightarrow> xs + vs \<le># ys + ws" | 
| 827 | apply (unfold mset_le_def) | |
| 828 | apply auto | |
| 15072 | 829 | apply (erule_tac x=a in allE)+ | 
| 830 | apply auto | |
| 831 | done | |
| 832 | ||
| 833 | lemma mset_le_add_left[simp]: "xs \<le># xs + ys" | |
| 18730 | 834 | unfolding mset_le_def by auto | 
| 15072 | 835 | |
| 836 | lemma mset_le_add_right[simp]: "ys \<le># xs + ys" | |
| 18730 | 837 | unfolding mset_le_def by auto | 
| 15072 | 838 | |
| 839 | lemma multiset_of_remdups_le: "multiset_of (remdups x) \<le># multiset_of x" | |
| 17200 | 840 | apply (induct x) | 
| 841 | apply auto | |
| 842 | apply (rule mset_le_trans) | |
| 843 | apply auto | |
| 844 | done | |
| 15072 | 845 | |
| 10249 | 846 | end |