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