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