| author | wenzelm | 
| Fri, 07 Dec 2007 22:19:45 +0100 | |
| changeset 25577 | d739f48ef40c | 
| parent 24893 | b8ef7afe3a6b | 
| child 26417 | af821e3a99e1 | 
| permissions | -rw-r--r-- | 
| 12610 | 1  | 
(* Title: ZF/Induct/Multiset.thy  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author: Sidi O Ehmety, Cambridge University Computer Laboratory  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
5  | 
A definitional theory of multisets,  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
6  | 
including a wellfoundedness proof for the multiset order.  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
8  | 
The theory features ordinal multisets and the usual ordering.  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
9  | 
*)  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
10  | 
|
| 15201 | 11  | 
theory Multiset  | 
12  | 
imports FoldSet Acc  | 
|
13  | 
begin  | 
|
14  | 
||
| 24892 | 15  | 
abbreviation (input)  | 
16  | 
  -- {* Short cut for multiset space *}
 | 
|
17  | 
Mult :: "i=>i" where  | 
|
18  | 
  "Mult(A) == A -||> nat-{0}"
 | 
|
| 15201 | 19  | 
|
| 24893 | 20  | 
definition  | 
| 12891 | 21  | 
(* This is the original "restrict" from ZF.thy.  | 
| 15201 | 22  | 
Restricts the function f to the domain A  | 
| 12891 | 23  | 
FIXME: adapt Multiset to the new "restrict". *)  | 
| 24893 | 24  | 
funrestrict :: "[i,i] => i" where  | 
| 15201 | 25  | 
"funrestrict(f,A) == \<lambda>x \<in> A. f`x"  | 
| 12891 | 26  | 
|
| 24893 | 27  | 
definition  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
28  | 
(* M is a multiset *)  | 
| 24893 | 29  | 
multiset :: "i => o" where  | 
| 15201 | 30  | 
  "multiset(M) == \<exists>A. M \<in> A -> nat-{0} & Finite(A)"
 | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
31  | 
|
| 24893 | 32  | 
definition  | 
33  | 
mset_of :: "i=>i" where  | 
|
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
34  | 
"mset_of(M) == domain(M)"  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
35  | 
|
| 24893 | 36  | 
definition  | 
37  | 
munion :: "[i, i] => i" (infixl "+#" 65) where  | 
|
| 15201 | 38  | 
"M +# N == \<lambda>x \<in> mset_of(M) Un mset_of(N).  | 
39  | 
if x \<in> mset_of(M) Int mset_of(N) then (M`x) #+ (N`x)  | 
|
40  | 
else (if x \<in> mset_of(M) then M`x else N`x)"  | 
|
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
41  | 
|
| 24893 | 42  | 
definition  | 
| 14046 | 43  | 
(*convert a function to a multiset by eliminating 0*)  | 
| 24893 | 44  | 
normalize :: "i => i" where  | 
| 14046 | 45  | 
"normalize(f) ==  | 
| 15201 | 46  | 
if (\<exists>A. f \<in> A -> nat & Finite(A)) then  | 
47  | 
            funrestrict(f, {x \<in> mset_of(f). 0 < f`x})
 | 
|
| 14046 | 48  | 
else 0"  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
49  | 
|
| 24893 | 50  | 
definition  | 
51  | 
mdiff :: "[i, i] => i" (infixl "-#" 65) where  | 
|
| 15201 | 52  | 
"M -# N == normalize(\<lambda>x \<in> mset_of(M).  | 
53  | 
if x \<in> mset_of(N) then M`x #- N`x else M`x)"  | 
|
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
54  | 
|
| 24893 | 55  | 
definition  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
56  | 
(* set of elements of a multiset *)  | 
| 24893 | 57  | 
  msingle :: "i => i"    ("{#_#}")  where
 | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
58  | 
  "{#a#} == {<a, 1>}"
 | 
| 15201 | 59  | 
|
| 24893 | 60  | 
definition  | 
61  | 
MCollect :: "[i, i=>o] => i" (*comprehension*) where  | 
|
| 15201 | 62  | 
  "MCollect(M, P) == funrestrict(M, {x \<in> mset_of(M). P(x)})"
 | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
63  | 
|
| 24893 | 64  | 
definition  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
65  | 
(* Counts the number of occurences of an element in a multiset *)  | 
| 24893 | 66  | 
mcount :: "[i, i] => i" where  | 
| 15201 | 67  | 
"mcount(M, a) == if a \<in> mset_of(M) then M`a else 0"  | 
68  | 
||
| 24893 | 69  | 
definition  | 
70  | 
msize :: "i => i" where  | 
|
| 15201 | 71  | 
"msize(M) == setsum(%a. $# mcount(M,a), mset_of(M))"  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
72  | 
|
| 24892 | 73  | 
abbreviation  | 
74  | 
  melem :: "[i,i] => o"    ("(_/ :# _)" [50, 51] 50)  where
 | 
|
75  | 
"a :# M == a \<in> mset_of(M)"  | 
|
76  | 
||
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
77  | 
syntax  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
78  | 
  "@MColl" :: "[pttrn, i, o] => i" ("(1{# _ : _./ _#})")
 | 
| 15201 | 79  | 
syntax (xsymbols)  | 
80  | 
  "@MColl" :: "[pttrn, i, o] => i" ("(1{# _ \<in> _./ _#})")
 | 
|
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
81  | 
translations  | 
| 24893 | 82  | 
  "{#x \<in> M. P#}" == "CONST MCollect(M, %x. P)"
 | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
83  | 
|
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
84  | 
(* multiset orderings *)  | 
| 15201 | 85  | 
|
| 24893 | 86  | 
definition  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
87  | 
(* multirel1 has to be a set (not a predicate) so that we can form  | 
| 
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
88  | 
its transitive closure and reason about wf(.) and acc(.) *)  | 
| 24893 | 89  | 
multirel1 :: "[i,i]=>i" where  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
90  | 
"multirel1(A, r) ==  | 
| 15201 | 91  | 
     {<M, N> \<in> Mult(A)*Mult(A).
 | 
92  | 
\<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A).  | 
|
93  | 
      N=M0 +# {#a#} & M=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}"
 | 
|
94  | 
||
| 24893 | 95  | 
definition  | 
96  | 
multirel :: "[i, i] => i" where  | 
|
| 15201 | 97  | 
"multirel(A, r) == multirel1(A, r)^+"  | 
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
98  | 
|
| 
12860
 
7fc3fbfed8ef
  Multiset:  added the translation Mult(A) => A-||>nat-{0}
 
paulson 
parents: 
12610 
diff
changeset
 | 
99  | 
(* ordinal multiset orderings *)  | 
| 15201 | 100  | 
|
| 24893 | 101  | 
definition  | 
102  | 
omultiset :: "i => o" where  | 
|
| 15201 | 103  | 
"omultiset(M) == \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))"  | 
104  | 
||
| 24893 | 105  | 
definition  | 
106  | 
mless :: "[i, i] => o" (infixl "<#" 50) where  | 
|
| 15201 | 107  | 
"M <# N == \<exists>i. Ord(i) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"  | 
108  | 
||
| 24893 | 109  | 
definition  | 
110  | 
mle :: "[i, i] => o" (infixl "<#=" 50) where  | 
|
| 15201 | 111  | 
"M <#= N == (omultiset(M) & M = N) | M <# N"  | 
112  | 
||
113  | 
||
114  | 
subsection{*Properties of the original "restrict" from ZF.thy*}
 | 
|
115  | 
||
116  | 
lemma funrestrict_subset: "[| f \<in> Pi(C,B); A\<subseteq>C |] ==> funrestrict(f,A) \<subseteq> f"  | 
|
117  | 
by (auto simp add: funrestrict_def lam_def intro: apply_Pair)  | 
|
118  | 
||
119  | 
lemma funrestrict_type:  | 
|
120  | 
"[| !!x. x \<in> A ==> f`x \<in> B(x) |] ==> funrestrict(f,A) \<in> Pi(A,B)"  | 
|
121  | 
by (simp add: funrestrict_def lam_type)  | 
|
122  | 
||
123  | 
lemma funrestrict_type2: "[| f \<in> Pi(C,B); A\<subseteq>C |] ==> funrestrict(f,A) \<in> Pi(A,B)"  | 
|
124  | 
by (blast intro: apply_type funrestrict_type)  | 
|
125  | 
||
126  | 
lemma funrestrict [simp]: "a \<in> A ==> funrestrict(f,A) ` a = f`a"  | 
|
127  | 
by (simp add: funrestrict_def)  | 
|
128  | 
||
129  | 
lemma funrestrict_empty [simp]: "funrestrict(f,0) = 0"  | 
|
130  | 
by (simp add: funrestrict_def)  | 
|
131  | 
||
132  | 
lemma domain_funrestrict [simp]: "domain(funrestrict(f,C)) = C"  | 
|
133  | 
by (auto simp add: funrestrict_def lam_def)  | 
|
134  | 
||
135  | 
lemma fun_cons_funrestrict_eq:  | 
|
136  | 
"f \<in> cons(a, b) -> B ==> f = cons(<a, f ` a>, funrestrict(f, b))"  | 
|
137  | 
apply (rule equalityI)  | 
|
138  | 
prefer 2 apply (blast intro: apply_Pair funrestrict_subset [THEN subsetD])  | 
|
139  | 
apply (auto dest!: Pi_memberD simp add: funrestrict_def lam_def)  | 
|
140  | 
done  | 
|
141  | 
||
142  | 
declare domain_of_fun [simp]  | 
|
143  | 
declare domainE [rule del]  | 
|
144  | 
||
145  | 
||
146  | 
text{* A useful simplification rule *}
 | 
|
147  | 
lemma multiset_fun_iff:  | 
|
148  | 
     "(f \<in> A -> nat-{0}) <-> f \<in> A->nat&(\<forall>a \<in> A. f`a \<in> nat & 0 < f`a)"
 | 
|
149  | 
apply safe  | 
|
150  | 
apply (rule_tac [4] B1 = "range (f) " in Pi_mono [THEN subsetD])  | 
|
151  | 
apply (auto intro!: Ord_0_lt  | 
|
152  | 
dest: apply_type Diff_subset [THEN Pi_mono, THEN subsetD]  | 
|
153  | 
simp add: range_of_fun apply_iff)  | 
|
154  | 
done  | 
|
155  | 
||
156  | 
(** The multiset space **)  | 
|
157  | 
lemma multiset_into_Mult: "[| multiset(M); mset_of(M)\<subseteq>A |] ==> M \<in> Mult(A)"  | 
|
158  | 
apply (simp add: multiset_def)  | 
|
159  | 
apply (auto simp add: multiset_fun_iff mset_of_def)  | 
|
160  | 
apply (rule_tac B1 = "nat-{0}" in FiniteFun_mono [THEN subsetD], simp_all)
 | 
|
161  | 
apply (rule Finite_into_Fin [THEN [2] Fin_mono [THEN subsetD], THEN fun_FiniteFunI])  | 
|
162  | 
apply (simp_all (no_asm_simp) add: multiset_fun_iff)  | 
|
163  | 
done  | 
|
164  | 
||
165  | 
lemma Mult_into_multiset: "M \<in> Mult(A) ==> multiset(M) & mset_of(M)\<subseteq>A"  | 
|
166  | 
apply (simp add: multiset_def mset_of_def)  | 
|
167  | 
apply (frule FiniteFun_is_fun)  | 
|
168  | 
apply (drule FiniteFun_domain_Fin)  | 
|
169  | 
apply (frule FinD, clarify)  | 
|
170  | 
apply (rule_tac x = "domain (M) " in exI)  | 
|
171  | 
apply (blast intro: Fin_into_Finite)  | 
|
172  | 
done  | 
|
173  | 
||
174  | 
lemma Mult_iff_multiset: "M \<in> Mult(A) <-> multiset(M) & mset_of(M)\<subseteq>A"  | 
|
175  | 
by (blast dest: Mult_into_multiset intro: multiset_into_Mult)  | 
|
176  | 
||
177  | 
lemma multiset_iff_Mult_mset_of: "multiset(M) <-> M \<in> Mult(mset_of(M))"  | 
|
178  | 
by (auto simp add: Mult_iff_multiset)  | 
|
179  | 
||
180  | 
||
181  | 
text{*The @{term multiset} operator*}
 | 
|
182  | 
||
183  | 
(* the empty multiset is 0 *)  | 
|
184  | 
||
185  | 
lemma multiset_0 [simp]: "multiset(0)"  | 
|
186  | 
by (auto intro: FiniteFun.intros simp add: multiset_iff_Mult_mset_of)  | 
|
187  | 
||
188  | 
||
189  | 
text{*The @{term mset_of} operator*}
 | 
|
190  | 
||
191  | 
lemma multiset_set_of_Finite [simp]: "multiset(M) ==> Finite(mset_of(M))"  | 
|
192  | 
by (simp add: multiset_def mset_of_def, auto)  | 
|
193  | 
||
194  | 
lemma mset_of_0 [iff]: "mset_of(0) = 0"  | 
|
195  | 
by (simp add: mset_of_def)  | 
|
196  | 
||
197  | 
lemma mset_is_0_iff: "multiset(M) ==> mset_of(M)=0 <-> M=0"  | 
|
198  | 
by (auto simp add: multiset_def mset_of_def)  | 
|
199  | 
||
200  | 
lemma mset_of_single [iff]: "mset_of({#a#}) = {a}"
 | 
|
201  | 
by (simp add: msingle_def mset_of_def)  | 
|
202  | 
||
203  | 
lemma mset_of_union [iff]: "mset_of(M +# N) = mset_of(M) Un mset_of(N)"  | 
|
204  | 
by (simp add: mset_of_def munion_def)  | 
|
205  | 
||
206  | 
lemma mset_of_diff [simp]: "mset_of(M)\<subseteq>A ==> mset_of(M -# N) \<subseteq> A"  | 
|
207  | 
by (auto simp add: mdiff_def multiset_def normalize_def mset_of_def)  | 
|
208  | 
||
209  | 
(* msingle *)  | 
|
210  | 
||
211  | 
lemma msingle_not_0 [iff]: "{#a#} \<noteq> 0 & 0 \<noteq> {#a#}"
 | 
|
212  | 
by (simp add: msingle_def)  | 
|
213  | 
||
214  | 
lemma msingle_eq_iff [iff]: "({#a#} = {#b#}) <->  (a = b)"
 | 
|
215  | 
by (simp add: msingle_def)  | 
|
216  | 
||
217  | 
lemma msingle_multiset [iff,TC]: "multiset({#a#})"
 | 
|
218  | 
apply (simp add: multiset_def msingle_def)  | 
|
219  | 
apply (rule_tac x = "{a}" in exI)
 | 
|
220  | 
apply (auto intro: Finite_cons Finite_0 fun_extend3)  | 
|
221  | 
done  | 
|
222  | 
||
223  | 
(** normalize **)  | 
|
224  | 
||
225  | 
lemmas Collect_Finite = Collect_subset [THEN subset_Finite, standard]  | 
|
226  | 
||
227  | 
lemma normalize_idem [simp]: "normalize(normalize(f)) = normalize(f)"  | 
|
228  | 
apply (simp add: normalize_def funrestrict_def mset_of_def)  | 
|
229  | 
apply (case_tac "\<exists>A. f \<in> A -> nat & Finite (A) ")  | 
|
230  | 
apply clarify  | 
|
231  | 
apply (drule_tac x = "{x \<in> domain (f) . 0 < f ` x}" in spec)
 | 
|
232  | 
apply auto  | 
|
233  | 
apply (auto intro!: lam_type simp add: Collect_Finite)  | 
|
234  | 
done  | 
|
235  | 
||
236  | 
lemma normalize_multiset [simp]: "multiset(M) ==> normalize(M) = M"  | 
|
237  | 
by (auto simp add: multiset_def normalize_def mset_of_def funrestrict_def multiset_fun_iff)  | 
|
238  | 
||
239  | 
lemma multiset_normalize [simp]: "multiset(normalize(f))"  | 
|
240  | 
apply (simp add: normalize_def)  | 
|
241  | 
apply (simp add: normalize_def mset_of_def multiset_def, auto)  | 
|
242  | 
apply (rule_tac x = "{x \<in> A . 0<f`x}" in exI)
 | 
|
243  | 
apply (auto intro: Collect_subset [THEN subset_Finite] funrestrict_type)  | 
|
244  | 
done  | 
|
245  | 
||
246  | 
(** Typechecking rules for union and difference of multisets **)  | 
|
247  | 
||
248  | 
(* union *)  | 
|
249  | 
||
250  | 
lemma munion_multiset [simp]: "[| multiset(M); multiset(N) |] ==> multiset(M +# N)"  | 
|
251  | 
apply (unfold multiset_def munion_def mset_of_def, auto)  | 
|
252  | 
apply (rule_tac x = "A Un Aa" in exI)  | 
|
253  | 
apply (auto intro!: lam_type intro: Finite_Un simp add: multiset_fun_iff zero_less_add)  | 
|
254  | 
done  | 
|
255  | 
||
256  | 
(* difference *)  | 
|
257  | 
||
258  | 
lemma mdiff_multiset [simp]: "multiset(M -# N)"  | 
|
259  | 
by (simp add: mdiff_def)  | 
|
260  | 
||
261  | 
(** Algebraic properties of multisets **)  | 
|
262  | 
||
263  | 
(* Union *)  | 
|
264  | 
||
265  | 
lemma munion_0 [simp]: "multiset(M) ==> M +# 0 = M & 0 +# M = M"  | 
|
266  | 
apply (simp add: multiset_def)  | 
|
267  | 
apply (auto simp add: munion_def mset_of_def)  | 
|
268  | 
done  | 
|
269  | 
||
270  | 
lemma munion_commute: "M +# N = N +# M"  | 
|
271  | 
by (auto intro!: lam_cong simp add: munion_def)  | 
|
272  | 
||
273  | 
lemma munion_assoc: "(M +# N) +# K = M +# (N +# K)"  | 
|
274  | 
apply (unfold munion_def mset_of_def)  | 
|
275  | 
apply (rule lam_cong, auto)  | 
|
276  | 
done  | 
|
277  | 
||
278  | 
lemma munion_lcommute: "M +# (N +# K) = N +# (M +# K)"  | 
|
279  | 
apply (unfold munion_def mset_of_def)  | 
|
280  | 
apply (rule lam_cong, auto)  | 
|
281  | 
done  | 
|
282  | 
||
283  | 
lemmas munion_ac = munion_commute munion_assoc munion_lcommute  | 
|
284  | 
||
285  | 
(* Difference *)  | 
|
286  | 
||
287  | 
lemma mdiff_self_eq_0 [simp]: "M -# M = 0"  | 
|
288  | 
by (simp add: mdiff_def normalize_def mset_of_def)  | 
|
289  | 
||
290  | 
lemma mdiff_0 [simp]: "0 -# M = 0"  | 
|
291  | 
by (simp add: mdiff_def normalize_def)  | 
|
292  | 
||
293  | 
lemma mdiff_0_right [simp]: "multiset(M) ==> M -# 0 = M"  | 
|
294  | 
by (auto simp add: multiset_def mdiff_def normalize_def multiset_fun_iff mset_of_def funrestrict_def)  | 
|
295  | 
||
296  | 
lemma mdiff_union_inverse2 [simp]: "multiset(M) ==> M +# {#a#} -# {#a#} = M"
 | 
|
297  | 
apply (unfold multiset_def munion_def mdiff_def msingle_def normalize_def mset_of_def)  | 
|
298  | 
apply (auto cong add: if_cong simp add: ltD multiset_fun_iff funrestrict_def subset_Un_iff2 [THEN iffD1])  | 
|
299  | 
prefer 2 apply (force intro!: lam_type)  | 
|
300  | 
apply (subgoal_tac [2] "{x \<in> A \<union> {a} . x \<noteq> a \<and> x \<in> A} = A")
 | 
|
301  | 
apply (rule fun_extension, auto)  | 
|
302  | 
apply (drule_tac x = "A Un {a}" in spec)
 | 
|
303  | 
apply (simp add: Finite_Un)  | 
|
304  | 
apply (force intro!: lam_type)  | 
|
305  | 
done  | 
|
306  | 
||
307  | 
(** Count of elements **)  | 
|
308  | 
||
309  | 
lemma mcount_type [simp,TC]: "multiset(M) ==> mcount(M, a) \<in> nat"  | 
|
310  | 
by (auto simp add: multiset_def mcount_def mset_of_def multiset_fun_iff)  | 
|
311  | 
||
312  | 
lemma mcount_0 [simp]: "mcount(0, a) = 0"  | 
|
313  | 
by (simp add: mcount_def)  | 
|
314  | 
||
315  | 
lemma mcount_single [simp]: "mcount({#b#}, a) = (if a=b then 1 else 0)"
 | 
|
316  | 
by (simp add: mcount_def mset_of_def msingle_def)  | 
|
317  | 
||
318  | 
lemma mcount_union [simp]: "[| multiset(M); multiset(N) |]  | 
|
319  | 
==> mcount(M +# N, a) = mcount(M, a) #+ mcount (N, a)"  | 
|
320  | 
apply (auto simp add: multiset_def multiset_fun_iff mcount_def munion_def mset_of_def)  | 
|
321  | 
done  | 
|
322  | 
||
323  | 
lemma mcount_diff [simp]:  | 
|
324  | 
"multiset(M) ==> mcount(M -# N, a) = mcount(M, a) #- mcount(N, a)"  | 
|
325  | 
apply (simp add: multiset_def)  | 
|
326  | 
apply (auto dest!: not_lt_imp_le  | 
|
327  | 
simp add: mdiff_def multiset_fun_iff mcount_def normalize_def mset_of_def)  | 
|
328  | 
apply (force intro!: lam_type)  | 
|
329  | 
apply (force intro!: lam_type)  | 
|
330  | 
done  | 
|
331  | 
||
332  | 
lemma mcount_elem: "[| multiset(M); a \<in> mset_of(M) |] ==> 0 < mcount(M, a)"  | 
|
333  | 
apply (simp add: multiset_def, clarify)  | 
|
334  | 
apply (simp add: mcount_def mset_of_def)  | 
|
335  | 
apply (simp add: multiset_fun_iff)  | 
|
336  | 
done  | 
|
337  | 
||
338  | 
(** msize **)  | 
|
339  | 
||
340  | 
lemma msize_0 [simp]: "msize(0) = #0"  | 
|
341  | 
by (simp add: msize_def)  | 
|
342  | 
||
343  | 
lemma msize_single [simp]: "msize({#a#}) = #1"
 | 
|
344  | 
by (simp add: msize_def)  | 
|
345  | 
||
346  | 
lemma msize_type [simp,TC]: "msize(M) \<in> int"  | 
|
347  | 
by (simp add: msize_def)  | 
|
348  | 
||
349  | 
lemma msize_zpositive: "multiset(M)==> #0 $\<le> msize(M)"  | 
|
350  | 
by (auto simp add: msize_def intro: g_zpos_imp_setsum_zpos)  | 
|
351  | 
||
352  | 
lemma msize_int_of_nat: "multiset(M) ==> \<exists>n \<in> nat. msize(M)= $# n"  | 
|
353  | 
apply (rule not_zneg_int_of)  | 
|
354  | 
apply (simp_all (no_asm_simp) add: msize_type [THEN znegative_iff_zless_0] not_zless_iff_zle msize_zpositive)  | 
|
355  | 
done  | 
|
356  | 
||
357  | 
lemma not_empty_multiset_imp_exist:  | 
|
358  | 
"[| M\<noteq>0; multiset(M) |] ==> \<exists>a \<in> mset_of(M). 0 < mcount(M, a)"  | 
|
359  | 
apply (simp add: multiset_def)  | 
|
360  | 
apply (erule not_emptyE)  | 
|
361  | 
apply (auto simp add: mset_of_def mcount_def multiset_fun_iff)  | 
|
362  | 
apply (blast dest!: fun_is_rel)  | 
|
363  | 
done  | 
|
364  | 
||
365  | 
lemma msize_eq_0_iff: "multiset(M) ==> msize(M)=#0 <-> M=0"  | 
|
366  | 
apply (simp add: msize_def, auto)  | 
|
367  | 
apply (rule_tac Pa = "setsum (?u,?v) \<noteq> #0" in swap)  | 
|
368  | 
apply blast  | 
|
369  | 
apply (drule not_empty_multiset_imp_exist, assumption, clarify)  | 
|
370  | 
apply (subgoal_tac "Finite (mset_of (M) - {a}) ")
 | 
|
371  | 
prefer 2 apply (simp add: Finite_Diff)  | 
|
372  | 
apply (subgoal_tac "setsum (%x. $# mcount (M, x), cons (a, mset_of (M) -{a}))=#0")
 | 
|
373  | 
prefer 2 apply (simp add: cons_Diff, simp)  | 
|
374  | 
apply (subgoal_tac "#0 $\<le> setsum (%x. $# mcount (M, x), mset_of (M) - {a}) ")
 | 
|
375  | 
apply (rule_tac [2] g_zpos_imp_setsum_zpos)  | 
|
376  | 
apply (auto simp add: Finite_Diff not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym])  | 
|
377  | 
apply (rule not_zneg_int_of [THEN bexE])  | 
|
378  | 
apply (auto simp del: int_of_0 simp add: int_of_add [symmetric] int_of_0 [symmetric])  | 
|
379  | 
done  | 
|
380  | 
||
381  | 
lemma setsum_mcount_Int:  | 
|
382  | 
"Finite(A) ==> setsum(%a. $# mcount(N, a), A Int mset_of(N))  | 
|
383  | 
= setsum(%a. $# mcount(N, a), A)"  | 
|
| 18415 | 384  | 
apply (induct rule: Finite_induct)  | 
385  | 
apply auto  | 
|
| 15201 | 386  | 
apply (subgoal_tac "Finite (B Int mset_of (N))")  | 
387  | 
prefer 2 apply (blast intro: subset_Finite)  | 
|
388  | 
apply (auto simp add: mcount_def Int_cons_left)  | 
|
389  | 
done  | 
|
390  | 
||
391  | 
lemma msize_union [simp]:  | 
|
392  | 
"[| multiset(M); multiset(N) |] ==> msize(M +# N) = msize(M) $+ msize(N)"  | 
|
393  | 
apply (simp add: msize_def setsum_Un setsum_addf int_of_add setsum_mcount_Int)  | 
|
394  | 
apply (subst Int_commute)  | 
|
395  | 
apply (simp add: setsum_mcount_Int)  | 
|
396  | 
done  | 
|
397  | 
||
398  | 
lemma msize_eq_succ_imp_elem: "[|msize(M)= $# succ(n); n \<in> nat|] ==> \<exists>a. a \<in> mset_of(M)"  | 
|
399  | 
apply (unfold msize_def)  | 
|
400  | 
apply (blast dest: setsum_succD)  | 
|
401  | 
done  | 
|
402  | 
||
403  | 
(** Equality of multisets **)  | 
|
404  | 
||
405  | 
lemma equality_lemma:  | 
|
406  | 
"[| multiset(M); multiset(N); \<forall>a. mcount(M, a)=mcount(N, a) |]  | 
|
407  | 
==> mset_of(M)=mset_of(N)"  | 
|
408  | 
apply (simp add: multiset_def)  | 
|
409  | 
apply (rule sym, rule equalityI)  | 
|
410  | 
apply (auto simp add: multiset_fun_iff mcount_def mset_of_def)  | 
|
411  | 
apply (drule_tac [!] x=x in spec)  | 
|
412  | 
apply (case_tac [2] "x \<in> Aa", case_tac "x \<in> A", auto)  | 
|
413  | 
done  | 
|
414  | 
||
415  | 
lemma multiset_equality:  | 
|
416  | 
"[| multiset(M); multiset(N) |]==> M=N<->(\<forall>a. mcount(M, a)=mcount(N, a))"  | 
|
417  | 
apply auto  | 
|
418  | 
apply (subgoal_tac "mset_of (M) = mset_of (N) ")  | 
|
419  | 
prefer 2 apply (blast intro: equality_lemma)  | 
|
420  | 
apply (simp add: multiset_def mset_of_def)  | 
|
421  | 
apply (auto simp add: multiset_fun_iff)  | 
|
422  | 
apply (rule fun_extension)  | 
|
423  | 
apply (blast, blast)  | 
|
424  | 
apply (drule_tac x = x in spec)  | 
|
425  | 
apply (auto simp add: mcount_def mset_of_def)  | 
|
426  | 
done  | 
|
427  | 
||
428  | 
(** More algebraic properties of multisets **)  | 
|
429  | 
||
430  | 
lemma munion_eq_0_iff [simp]: "[|multiset(M); multiset(N)|]==>(M +# N =0) <-> (M=0 & N=0)"  | 
|
431  | 
by (auto simp add: multiset_equality)  | 
|
432  | 
||
433  | 
lemma empty_eq_munion_iff [simp]: "[|multiset(M); multiset(N)|]==>(0=M +# N) <-> (M=0 & N=0)"  | 
|
434  | 
apply (rule iffI, drule sym)  | 
|
435  | 
apply (simp_all add: multiset_equality)  | 
|
436  | 
done  | 
|
437  | 
||
438  | 
lemma munion_right_cancel [simp]:  | 
|
439  | 
"[| multiset(M); multiset(N); multiset(K) |]==>(M +# K = N +# K)<->(M=N)"  | 
|
440  | 
by (auto simp add: multiset_equality)  | 
|
441  | 
||
442  | 
lemma munion_left_cancel [simp]:  | 
|
443  | 
"[|multiset(K); multiset(M); multiset(N)|] ==>(K +# M = K +# N) <-> (M = N)"  | 
|
444  | 
by (auto simp add: multiset_equality)  | 
|
445  | 
||
446  | 
lemma nat_add_eq_1_cases: "[| m \<in> nat; n \<in> nat |] ==> (m #+ n = 1) <-> (m=1 & n=0) | (m=0 & n=1)"  | 
|
| 18415 | 447  | 
by (induct_tac n) auto  | 
| 15201 | 448  | 
|
449  | 
lemma munion_is_single:  | 
|
450  | 
"[|multiset(M); multiset(N)|]  | 
|
451  | 
      ==> (M +# N = {#a#}) <->  (M={#a#} & N=0) | (M = 0 & N = {#a#})"
 | 
|
452  | 
apply (simp (no_asm_simp) add: multiset_equality)  | 
|
453  | 
apply safe  | 
|
454  | 
apply simp_all  | 
|
455  | 
apply (case_tac "aa=a")  | 
|
456  | 
apply (drule_tac [2] x = aa in spec)  | 
|
457  | 
apply (drule_tac x = a in spec)  | 
|
458  | 
apply (simp add: nat_add_eq_1_cases, simp)  | 
|
459  | 
apply (case_tac "aaa=aa", simp)  | 
|
460  | 
apply (drule_tac x = aa in spec)  | 
|
461  | 
apply (simp add: nat_add_eq_1_cases)  | 
|
462  | 
apply (case_tac "aaa=a")  | 
|
463  | 
apply (drule_tac [4] x = aa in spec)  | 
|
464  | 
apply (drule_tac [3] x = a in spec)  | 
|
465  | 
apply (drule_tac [2] x = aaa in spec)  | 
|
466  | 
apply (drule_tac x = aa in spec)  | 
|
467  | 
apply (simp_all add: nat_add_eq_1_cases)  | 
|
468  | 
done  | 
|
469  | 
||
470  | 
lemma msingle_is_union: "[| multiset(M); multiset(N) |]  | 
|
471  | 
  ==> ({#a#} = M +# N) <-> ({#a#} = M  & N=0 | M = 0 & {#a#} = N)"
 | 
|
472  | 
apply (subgoal_tac " ({#a#} = M +# N) <-> (M +# N = {#a#}) ")
 | 
|
473  | 
apply (simp (no_asm_simp) add: munion_is_single)  | 
|
474  | 
apply blast  | 
|
475  | 
apply (blast dest: sym)  | 
|
476  | 
done  | 
|
477  | 
||
478  | 
(** Towards induction over multisets **)  | 
|
479  | 
||
480  | 
lemma setsum_decr:  | 
|
481  | 
"Finite(A)  | 
|
482  | 
==> (\<forall>M. multiset(M) -->  | 
|
483  | 
(\<forall>a \<in> mset_of(M). setsum(%z. $# mcount(M(a:=M`a #- 1), z), A) =  | 
|
484  | 
(if a \<in> A then setsum(%z. $# mcount(M, z), A) $- #1  | 
|
485  | 
else setsum(%z. $# mcount(M, z), A))))"  | 
|
486  | 
apply (unfold multiset_def)  | 
|
487  | 
apply (erule Finite_induct)  | 
|
488  | 
apply (auto simp add: multiset_fun_iff)  | 
|
489  | 
apply (unfold mset_of_def mcount_def)  | 
|
490  | 
apply (case_tac "x \<in> A", auto)  | 
|
491  | 
apply (subgoal_tac "$# M ` x $+ #-1 = $# M ` x $- $# 1")  | 
|
492  | 
apply (erule ssubst)  | 
|
493  | 
apply (rule int_of_diff, auto)  | 
|
494  | 
done  | 
|
495  | 
||
496  | 
lemma setsum_decr2:  | 
|
497  | 
"Finite(A)  | 
|
498  | 
==> \<forall>M. multiset(M) --> (\<forall>a \<in> mset_of(M).  | 
|
| 16973 | 499  | 
           setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) =
 | 
500  | 
(if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a  | 
|
501  | 
else setsum(%x. $# mcount(M, x), A)))"  | 
|
| 15201 | 502  | 
apply (simp add: multiset_def)  | 
503  | 
apply (erule Finite_induct)  | 
|
504  | 
apply (auto simp add: multiset_fun_iff mcount_def mset_of_def)  | 
|
505  | 
done  | 
|
506  | 
||
507  | 
lemma setsum_decr3: "[| Finite(A); multiset(M); a \<in> mset_of(M) |]  | 
|
508  | 
      ==> setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) =
 | 
|
509  | 
(if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a  | 
|
510  | 
else setsum(%x. $# mcount(M, x), A))"  | 
|
511  | 
apply (subgoal_tac "setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ")
 | 
|
512  | 
apply (rule_tac [2] setsum_Diff [symmetric])  | 
|
513  | 
apply (rule sym, rule ssubst, blast)  | 
|
514  | 
apply (rule sym, drule setsum_decr2, auto)  | 
|
515  | 
apply (simp add: mcount_def mset_of_def)  | 
|
516  | 
done  | 
|
517  | 
||
518  | 
lemma nat_le_1_cases: "n \<in> nat ==> n le 1 <-> (n=0 | n=1)"  | 
|
519  | 
by (auto elim: natE)  | 
|
520  | 
||
521  | 
lemma succ_pred_eq_self: "[| 0<n; n \<in> nat |] ==> succ(n #- 1) = n"  | 
|
522  | 
apply (subgoal_tac "1 le n")  | 
|
523  | 
apply (drule add_diff_inverse2, auto)  | 
|
524  | 
done  | 
|
525  | 
||
526  | 
text{*Specialized for use in the proof below.*}
 | 
|
527  | 
lemma multiset_funrestict:  | 
|
528  | 
"\<lbrakk>\<forall>a\<in>A. M ` a \<in> nat \<and> 0 < M ` a; Finite(A)\<rbrakk>  | 
|
529  | 
      \<Longrightarrow> multiset(funrestrict(M, A - {a}))"
 | 
|
530  | 
apply (simp add: multiset_def multiset_fun_iff)  | 
|
531  | 
apply (rule_tac x="A-{a}" in exI)
 | 
|
532  | 
apply (auto intro: Finite_Diff funrestrict_type)  | 
|
533  | 
done  | 
|
534  | 
||
535  | 
lemma multiset_induct_aux:  | 
|
536  | 
assumes prem1: "!!M a. [| multiset(M); a\<notin>mset_of(M); P(M) |] ==> P(cons(<a, 1>, M))"  | 
|
537  | 
and prem2: "!!M b. [| multiset(M); b \<in> mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))"  | 
|
538  | 
shows  | 
|
539  | 
"[| n \<in> nat; P(0) |]  | 
|
540  | 
==> (\<forall>M. multiset(M)-->  | 
|
541  | 
  (setsum(%x. $# mcount(M, x), {x \<in> mset_of(M). 0 < M`x}) = $# n) --> P(M))"
 | 
|
542  | 
apply (erule nat_induct, clarify)  | 
|
543  | 
apply (frule msize_eq_0_iff)  | 
|
544  | 
apply (auto simp add: mset_of_def multiset_def multiset_fun_iff msize_def)  | 
|
545  | 
apply (subgoal_tac "setsum (%x. $# mcount (M, x), A) =$# succ (x) ")  | 
|
546  | 
apply (drule setsum_succD, auto)  | 
|
547  | 
apply (case_tac "1 <M`a")  | 
|
548  | 
apply (drule_tac [2] not_lt_imp_le)  | 
|
549  | 
apply (simp_all add: nat_le_1_cases)  | 
|
550  | 
apply (subgoal_tac "M= (M (a:=M`a #- 1)) (a:= (M (a:=M`a #- 1))`a #+ 1) ")  | 
|
551  | 
apply (rule_tac [2] A = A and B = "%x. nat" and D = "%x. nat" in fun_extension)  | 
|
552  | 
apply (rule_tac [3] update_type)+  | 
|
553  | 
apply (simp_all (no_asm_simp))  | 
|
554  | 
apply (rule_tac [2] impI)  | 
|
555  | 
apply (rule_tac [2] succ_pred_eq_self [symmetric])  | 
|
556  | 
apply (simp_all (no_asm_simp))  | 
|
557  | 
apply (rule subst, rule sym, blast, rule prem2)  | 
|
558  | 
apply (simp (no_asm) add: multiset_def multiset_fun_iff)  | 
|
559  | 
apply (rule_tac x = A in exI)  | 
|
560  | 
apply (force intro: update_type)  | 
|
561  | 
apply (simp (no_asm_simp) add: mset_of_def mcount_def)  | 
|
562  | 
apply (drule_tac x = "M (a := M ` a #- 1) " in spec)  | 
|
563  | 
apply (drule mp, drule_tac [2] mp, simp_all)  | 
|
564  | 
apply (rule_tac x = A in exI)  | 
|
565  | 
apply (auto intro: update_type)  | 
|
566  | 
apply (subgoal_tac "Finite ({x \<in> cons (a, A) . x\<noteq>a-->0<M`x}) ")
 | 
|
567  | 
prefer 2 apply (blast intro: Collect_subset [THEN subset_Finite] Finite_cons)  | 
|
568  | 
apply (drule_tac A = "{x \<in> cons (a, A) . x\<noteq>a-->0<M`x}" in setsum_decr)
 | 
|
569  | 
apply (drule_tac x = M in spec)  | 
|
570  | 
apply (subgoal_tac "multiset (M) ")  | 
|
571  | 
prefer 2  | 
|
572  | 
apply (simp add: multiset_def multiset_fun_iff)  | 
|
573  | 
apply (rule_tac x = A in exI, force)  | 
|
574  | 
apply (simp_all add: mset_of_def)  | 
|
575  | 
apply (drule_tac psi = "\<forall>x \<in> A. ?u (x) " in asm_rl)  | 
|
576  | 
apply (drule_tac x = a in bspec)  | 
|
577  | 
apply (simp (no_asm_simp))  | 
|
578  | 
apply (subgoal_tac "cons (a, A) = A")  | 
|
579  | 
prefer 2 apply blast  | 
|
580  | 
apply simp  | 
|
581  | 
apply (subgoal_tac "M=cons (<a, M`a>, funrestrict (M, A-{a}))")
 | 
|
582  | 
prefer 2  | 
|
583  | 
apply (rule fun_cons_funrestrict_eq)  | 
|
584  | 
 apply (subgoal_tac "cons (a, A-{a}) = A")
 | 
|
585  | 
apply force  | 
|
586  | 
apply force  | 
|
587  | 
apply (rule_tac a = "cons (<a, 1>, funrestrict (M, A - {a}))" in ssubst)
 | 
|
588  | 
apply simp  | 
|
589  | 
apply (frule multiset_funrestict, assumption)  | 
|
590  | 
apply (rule prem1, assumption)  | 
|
591  | 
apply (simp add: mset_of_def)  | 
|
592  | 
apply (drule_tac x = "funrestrict (M, A-{a}) " in spec)
 | 
|
593  | 
apply (drule mp)  | 
|
594  | 
apply (rule_tac x = "A-{a}" in exI)
 | 
|
595  | 
apply (auto intro: Finite_Diff funrestrict_type simp add: funrestrict)  | 
|
596  | 
apply (frule_tac A = A and M = M and a = a in setsum_decr3)  | 
|
597  | 
apply (simp (no_asm_simp) add: multiset_def multiset_fun_iff)  | 
|
598  | 
apply blast  | 
|
599  | 
apply (simp (no_asm_simp) add: mset_of_def)  | 
|
600  | 
apply (drule_tac b = "if ?u then ?v else ?w" in sym, simp_all)  | 
|
601  | 
apply (subgoal_tac "{x \<in> A - {a} . 0 < funrestrict (M, A - {x}) ` x} = A - {a}")
 | 
|
602  | 
apply (auto intro!: setsum_cong simp add: zdiff_eq_iff zadd_commute multiset_def multiset_fun_iff mset_of_def)  | 
|
603  | 
done  | 
|
604  | 
||
605  | 
lemma multiset_induct2:  | 
|
606  | 
"[| multiset(M); P(0);  | 
|
607  | 
(!!M a. [| multiset(M); a\<notin>mset_of(M); P(M) |] ==> P(cons(<a, 1>, M)));  | 
|
608  | 
(!!M b. [| multiset(M); b \<in> mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))) |]  | 
|
609  | 
==> P(M)"  | 
|
610  | 
apply (subgoal_tac "\<exists>n \<in> nat. setsum (\<lambda>x. $# mcount (M, x), {x \<in> mset_of (M) . 0 < M ` x}) = $# n")
 | 
|
611  | 
apply (rule_tac [2] not_zneg_int_of)  | 
|
612  | 
apply (simp_all (no_asm_simp) add: znegative_iff_zless_0 not_zless_iff_zle)  | 
|
613  | 
apply (rule_tac [2] g_zpos_imp_setsum_zpos)  | 
|
614  | 
prefer 2 apply (blast intro: multiset_set_of_Finite Collect_subset [THEN subset_Finite])  | 
|
615  | 
prefer 2 apply (simp add: multiset_def multiset_fun_iff, clarify)  | 
|
616  | 
apply (rule multiset_induct_aux [rule_format], auto)  | 
|
617  | 
done  | 
|
618  | 
||
619  | 
lemma munion_single_case1:  | 
|
620  | 
     "[| multiset(M); a \<notin>mset_of(M) |] ==> M +# {#a#} = cons(<a, 1>, M)"
 | 
|
621  | 
apply (simp add: multiset_def msingle_def)  | 
|
622  | 
apply (auto simp add: munion_def)  | 
|
623  | 
apply (unfold mset_of_def, simp)  | 
|
624  | 
apply (rule fun_extension, rule lam_type, simp_all)  | 
|
625  | 
apply (auto simp add: multiset_fun_iff fun_extend_apply)  | 
|
626  | 
apply (drule_tac c = a and b = 1 in fun_extend3)  | 
|
627  | 
apply (auto simp add: cons_eq Un_commute [of _ "{a}"])
 | 
|
628  | 
done  | 
|
629  | 
||
630  | 
lemma munion_single_case2:  | 
|
631  | 
     "[| multiset(M); a \<in> mset_of(M) |] ==> M +# {#a#} = M(a:=M`a #+ 1)"
 | 
|
632  | 
apply (simp add: multiset_def)  | 
|
633  | 
apply (auto simp add: munion_def multiset_fun_iff msingle_def)  | 
|
634  | 
apply (unfold mset_of_def, simp)  | 
|
635  | 
apply (subgoal_tac "A Un {a} = A")
 | 
|
636  | 
apply (rule fun_extension)  | 
|
637  | 
apply (auto dest: domain_type intro: lam_type update_type)  | 
|
638  | 
done  | 
|
639  | 
||
640  | 
(* Induction principle for multisets *)  | 
|
641  | 
||
642  | 
lemma multiset_induct:  | 
|
643  | 
assumes M: "multiset(M)"  | 
|
644  | 
and P0: "P(0)"  | 
|
645  | 
      and step: "!!M a. [| multiset(M); P(M) |] ==> P(M +# {#a#})"
 | 
|
646  | 
shows "P(M)"  | 
|
647  | 
apply (rule multiset_induct2 [OF M])  | 
|
648  | 
apply (simp_all add: P0)  | 
|
| 20898 | 649  | 
apply (frule_tac [2] a = b in munion_single_case2 [symmetric])  | 
650  | 
apply (frule_tac a = a in munion_single_case1 [symmetric])  | 
|
| 15201 | 651  | 
apply (auto intro: step)  | 
652  | 
done  | 
|
653  | 
||
654  | 
(** MCollect **)  | 
|
655  | 
||
656  | 
lemma MCollect_multiset [simp]:  | 
|
657  | 
     "multiset(M) ==> multiset({# x \<in> M. P(x)#})"
 | 
|
658  | 
apply (simp add: MCollect_def multiset_def mset_of_def, clarify)  | 
|
659  | 
apply (rule_tac x = "{x \<in> A. P (x) }" in exI)
 | 
|
660  | 
apply (auto dest: CollectD1 [THEN [2] apply_type]  | 
|
661  | 
intro: Collect_subset [THEN subset_Finite] funrestrict_type)  | 
|
662  | 
done  | 
|
663  | 
||
664  | 
lemma mset_of_MCollect [simp]:  | 
|
665  | 
     "multiset(M) ==> mset_of({# x \<in> M. P(x) #}) \<subseteq> mset_of(M)"
 | 
|
666  | 
by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def)  | 
|
667  | 
||
668  | 
lemma MCollect_mem_iff [iff]:  | 
|
669  | 
     "x \<in> mset_of({#x \<in> M. P(x)#}) <->  x \<in> mset_of(M) & P(x)"
 | 
|
670  | 
by (simp add: MCollect_def mset_of_def)  | 
|
671  | 
||
672  | 
lemma mcount_MCollect [simp]:  | 
|
673  | 
     "mcount({# x \<in> M. P(x) #}, a) = (if P(a) then mcount(M,a) else 0)"
 | 
|
674  | 
by (simp add: mcount_def MCollect_def mset_of_def)  | 
|
675  | 
||
676  | 
lemma multiset_partition: "multiset(M) ==> M = {# x \<in> M. P(x) #} +# {# x \<in> M. ~ P(x) #}"
 | 
|
677  | 
by (simp add: multiset_equality)  | 
|
678  | 
||
679  | 
lemma natify_elem_is_self [simp]:  | 
|
680  | 
"[| multiset(M); a \<in> mset_of(M) |] ==> natify(M`a) = M`a"  | 
|
681  | 
by (auto simp add: multiset_def mset_of_def multiset_fun_iff)  | 
|
682  | 
||
683  | 
(* and more algebraic laws on multisets *)  | 
|
684  | 
||
685  | 
lemma munion_eq_conv_diff: "[| multiset(M); multiset(N) |]  | 
|
686  | 
  ==>  (M +# {#a#} = N +# {#b#}) <->  (M = N & a = b |
 | 
|
687  | 
       M = N -# {#a#} +# {#b#} & N = M -# {#b#} +# {#a#})"
 | 
|
688  | 
apply (simp del: mcount_single add: multiset_equality)  | 
|
689  | 
apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE)  | 
|
690  | 
apply (case_tac "a=b", auto)  | 
|
691  | 
apply (drule_tac x = a in spec)  | 
|
692  | 
apply (drule_tac [2] x = b in spec)  | 
|
693  | 
apply (drule_tac [3] x = aa in spec)  | 
|
694  | 
apply (drule_tac [4] x = a in spec, auto)  | 
|
695  | 
apply (subgoal_tac [!] "mcount (N,a) :nat")  | 
|
696  | 
apply (erule_tac [3] natE, erule natE, auto)  | 
|
697  | 
done  | 
|
698  | 
||
699  | 
lemma melem_diff_single:  | 
|
700  | 
"multiset(M) ==>  | 
|
701  | 
  k \<in> mset_of(M -# {#a#}) <-> (k=a & 1 < mcount(M,a)) | (k\<noteq> a & k \<in> mset_of(M))"
 | 
|
702  | 
apply (simp add: multiset_def)  | 
|
703  | 
apply (simp add: normalize_def mset_of_def msingle_def mdiff_def mcount_def)  | 
|
704  | 
apply (auto dest: domain_type intro: zero_less_diff [THEN iffD1]  | 
|
705  | 
simp add: multiset_fun_iff apply_iff)  | 
|
706  | 
apply (force intro!: lam_type)  | 
|
707  | 
apply (force intro!: lam_type)  | 
|
708  | 
apply (force intro!: lam_type)  | 
|
709  | 
done  | 
|
710  | 
||
711  | 
lemma munion_eq_conv_exist:  | 
|
712  | 
"[| M \<in> Mult(A); N \<in> Mult(A) |]  | 
|
713  | 
  ==> (M +# {#a#} = N +# {#b#}) <->
 | 
|
714  | 
      (M=N & a=b | (\<exists>K \<in> Mult(A). M= K +# {#b#} & N=K +# {#a#}))"
 | 
|
715  | 
by (auto simp add: Mult_iff_multiset melem_diff_single munion_eq_conv_diff)  | 
|
716  | 
||
717  | 
||
718  | 
subsection{*Multiset Orderings*}
 | 
|
719  | 
||
720  | 
(* multiset on a domain A are finite functions from A to nat-{0} *)
 | 
|
721  | 
||
722  | 
||
723  | 
(* multirel1 type *)  | 
|
724  | 
||
725  | 
lemma multirel1_type: "multirel1(A, r) \<subseteq> Mult(A)*Mult(A)"  | 
|
726  | 
by (auto simp add: multirel1_def)  | 
|
727  | 
||
728  | 
lemma multirel1_0 [simp]: "multirel1(0, r) =0"  | 
|
729  | 
by (auto simp add: multirel1_def)  | 
|
730  | 
||
731  | 
lemma multirel1_iff:  | 
|
732  | 
" <N, M> \<in> multirel1(A, r) <->  | 
|
733  | 
(\<exists>a. a \<in> A &  | 
|
734  | 
(\<exists>M0. M0 \<in> Mult(A) & (\<exists>K. K \<in> Mult(A) &  | 
|
735  | 
   M=M0 +# {#a#} & N=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r))))"
 | 
|
736  | 
by (auto simp add: multirel1_def Mult_iff_multiset Bex_def)  | 
|
737  | 
||
738  | 
||
739  | 
text{*Monotonicity of @{term multirel1}*}
 | 
|
740  | 
||
741  | 
lemma multirel1_mono1: "A\<subseteq>B ==> multirel1(A, r)\<subseteq>multirel1(B, r)"  | 
|
742  | 
apply (auto simp add: multirel1_def)  | 
|
743  | 
apply (auto simp add: Un_subset_iff Mult_iff_multiset)  | 
|
744  | 
apply (rule_tac x = a in bexI)  | 
|
745  | 
apply (rule_tac x = M0 in bexI, simp)  | 
|
746  | 
apply (rule_tac x = K in bexI)  | 
|
747  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
748  | 
done  | 
|
749  | 
||
750  | 
lemma multirel1_mono2: "r\<subseteq>s ==> multirel1(A,r)\<subseteq>multirel1(A, s)"  | 
|
751  | 
apply (simp add: multirel1_def, auto)  | 
|
752  | 
apply (rule_tac x = a in bexI)  | 
|
753  | 
apply (rule_tac x = M0 in bexI)  | 
|
754  | 
apply (simp_all add: Mult_iff_multiset)  | 
|
755  | 
apply (rule_tac x = K in bexI)  | 
|
756  | 
apply (simp_all add: Mult_iff_multiset, auto)  | 
|
757  | 
done  | 
|
758  | 
||
759  | 
lemma multirel1_mono:  | 
|
760  | 
"[| A\<subseteq>B; r\<subseteq>s |] ==> multirel1(A, r) \<subseteq> multirel1(B, s)"  | 
|
761  | 
apply (rule subset_trans)  | 
|
762  | 
apply (rule multirel1_mono1)  | 
|
763  | 
apply (rule_tac [2] multirel1_mono2, auto)  | 
|
764  | 
done  | 
|
765  | 
||
766  | 
subsection{* Toward the proof of well-foundedness of multirel1 *}
 | 
|
767  | 
||
768  | 
lemma not_less_0 [iff]: "<M,0> \<notin> multirel1(A, r)"  | 
|
769  | 
by (auto simp add: multirel1_def Mult_iff_multiset)  | 
|
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
770  | 
|
| 15201 | 771  | 
lemma less_munion: "[| <N, M0 +# {#a#}> \<in> multirel1(A, r); M0 \<in> Mult(A) |] ==>
 | 
772  | 
  (\<exists>M. <M, M0> \<in> multirel1(A, r) & N = M +# {#a#}) |
 | 
|
773  | 
(\<exists>K. K \<in> Mult(A) & (\<forall>b \<in> mset_of(K). <b, a> \<in> r) & N = M0 +# K)"  | 
|
774  | 
apply (frule multirel1_type [THEN subsetD])  | 
|
775  | 
apply (simp add: multirel1_iff)  | 
|
776  | 
apply (auto simp add: munion_eq_conv_exist)  | 
|
777  | 
apply (rule_tac x="Ka +# K" in exI, auto, simp add: Mult_iff_multiset)  | 
|
778  | 
apply (simp (no_asm_simp) add: munion_left_cancel munion_assoc)  | 
|
779  | 
apply (auto simp add: munion_commute)  | 
|
780  | 
done  | 
|
781  | 
||
782  | 
lemma multirel1_base: "[| M \<in> Mult(A); a \<in> A |] ==> <M, M +# {#a#}> \<in> multirel1(A, r)"
 | 
|
783  | 
apply (auto simp add: multirel1_iff)  | 
|
784  | 
apply (simp add: Mult_iff_multiset)  | 
|
785  | 
apply (rule_tac x = a in exI, clarify)  | 
|
786  | 
apply (rule_tac x = M in exI, simp)  | 
|
787  | 
apply (rule_tac x = 0 in exI, auto)  | 
|
788  | 
done  | 
|
789  | 
||
790  | 
lemma acc_0: "acc(0)=0"  | 
|
791  | 
by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD])  | 
|
792  | 
||
793  | 
lemma lemma1: "[| \<forall>b \<in> A. <b,a> \<in> r -->  | 
|
794  | 
    (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r)));
 | 
|
795  | 
M0 \<in> acc(multirel1(A, r)); a \<in> A;  | 
|
796  | 
    \<forall>M. <M,M0> \<in> multirel1(A, r) --> M +# {#a#} \<in> acc(multirel1(A, r)) |]
 | 
|
797  | 
  ==> M0 +# {#a#} \<in> acc(multirel1(A, r))"
 | 
|
| 15481 | 798  | 
apply (subgoal_tac "M0 \<in> Mult(A) ")  | 
| 15201 | 799  | 
prefer 2  | 
800  | 
apply (erule acc.cases)  | 
|
801  | 
apply (erule fieldE)  | 
|
802  | 
apply (auto dest: multirel1_type [THEN subsetD])  | 
|
803  | 
apply (rule accI)  | 
|
804  | 
apply (rename_tac "N")  | 
|
805  | 
apply (drule less_munion, blast)  | 
|
806  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
807  | 
apply (erule_tac P = "\<forall>x \<in> mset_of (K) . <x, a> \<in> r" in rev_mp)  | 
|
808  | 
apply (erule_tac P = "mset_of (K) \<subseteq>A" in rev_mp)  | 
|
809  | 
apply (erule_tac M = K in multiset_induct)  | 
|
810  | 
(* three subgoals *)  | 
|
811  | 
(* subgoal 1: the induction base case *)  | 
|
812  | 
apply (simp (no_asm_simp))  | 
|
813  | 
(* subgoal 2: the induction general case *)  | 
|
814  | 
apply (simp add: Ball_def Un_subset_iff, clarify)  | 
|
815  | 
apply (drule_tac x = aa in spec, simp)  | 
|
816  | 
apply (subgoal_tac "aa \<in> A")  | 
|
817  | 
prefer 2 apply blast  | 
|
818  | 
apply (drule_tac x = "M0 +# M" and P =  | 
|
819  | 
"%x. x \<in> acc(multirel1(A, r)) \<longrightarrow> ?Q(x)" in spec)  | 
|
820  | 
apply (simp add: munion_assoc [symmetric])  | 
|
821  | 
(* subgoal 3: additional conditions *)  | 
|
822  | 
apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset)  | 
|
823  | 
done  | 
|
824  | 
||
825  | 
lemma lemma2: "[| \<forall>b \<in> A. <b,a> \<in> r  | 
|
826  | 
   --> (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#} :acc(multirel1(A, r)));
 | 
|
827  | 
        M \<in> acc(multirel1(A, r)); a \<in> A|] ==> M +# {#a#} \<in> acc(multirel1(A, r))"
 | 
|
828  | 
apply (erule acc_induct)  | 
|
829  | 
apply (blast intro: lemma1)  | 
|
830  | 
done  | 
|
831  | 
||
832  | 
lemma lemma3: "[| wf[A](r); a \<in> A |]  | 
|
833  | 
      ==> \<forall>M \<in> acc(multirel1(A, r)). M +# {#a#} \<in> acc(multirel1(A, r))"
 | 
|
834  | 
apply (erule_tac a = a in wf_on_induct, blast)  | 
|
835  | 
apply (blast intro: lemma2)  | 
|
836  | 
done  | 
|
837  | 
||
838  | 
lemma lemma4: "multiset(M) ==> mset_of(M)\<subseteq>A -->  | 
|
839  | 
wf[A](r) --> M \<in> field(multirel1(A, r)) --> M \<in> acc(multirel1(A, r))"  | 
|
840  | 
apply (erule multiset_induct)  | 
|
841  | 
(* proving the base case *)  | 
|
842  | 
apply clarify  | 
|
843  | 
apply (rule accI, force)  | 
|
844  | 
apply (simp add: multirel1_def)  | 
|
845  | 
(* Proving the general case *)  | 
|
846  | 
apply clarify  | 
|
847  | 
apply simp  | 
|
848  | 
apply (subgoal_tac "mset_of (M) \<subseteq>A")  | 
|
849  | 
prefer 2 apply blast  | 
|
850  | 
apply clarify  | 
|
851  | 
apply (drule_tac a = a in lemma3, blast)  | 
|
852  | 
apply (subgoal_tac "M \<in> field (multirel1 (A,r))")  | 
|
853  | 
apply blast  | 
|
854  | 
apply (rule multirel1_base [THEN fieldI1])  | 
|
855  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
856  | 
done  | 
|
857  | 
||
858  | 
lemma all_accessible: "[| wf[A](r); M \<in> Mult(A); A \<noteq> 0|] ==> M \<in> acc(multirel1(A, r))"  | 
|
859  | 
apply (erule not_emptyE)  | 
|
860  | 
apply (rule lemma4 [THEN mp, THEN mp, THEN mp])  | 
|
861  | 
apply (rule_tac [4] multirel1_base [THEN fieldI1])  | 
|
862  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
863  | 
done  | 
|
864  | 
||
865  | 
lemma wf_on_multirel1: "wf[A](r) ==> wf[A-||>nat-{0}](multirel1(A, r))"
 | 
|
866  | 
apply (case_tac "A=0")  | 
|
867  | 
apply (simp (no_asm_simp))  | 
|
868  | 
apply (rule wf_imp_wf_on)  | 
|
869  | 
apply (rule wf_on_field_imp_wf)  | 
|
870  | 
apply (simp (no_asm_simp) add: wf_on_0)  | 
|
871  | 
apply (rule_tac A = "acc (multirel1 (A,r))" in wf_on_subset_A)  | 
|
872  | 
apply (rule wf_on_acc)  | 
|
873  | 
apply (blast intro: all_accessible)  | 
|
874  | 
done  | 
|
875  | 
||
876  | 
lemma wf_multirel1: "wf(r) ==>wf(multirel1(field(r), r))"  | 
|
877  | 
apply (simp (no_asm_use) add: wf_iff_wf_on_field)  | 
|
878  | 
apply (drule wf_on_multirel1)  | 
|
879  | 
apply (rule_tac A = "field (r) -||> nat - {0}" in wf_on_subset_A)
 | 
|
880  | 
apply (simp (no_asm_simp))  | 
|
881  | 
apply (rule field_rel_subset)  | 
|
882  | 
apply (rule multirel1_type)  | 
|
883  | 
done  | 
|
884  | 
||
885  | 
(** multirel **)  | 
|
886  | 
||
887  | 
lemma multirel_type: "multirel(A, r) \<subseteq> Mult(A)*Mult(A)"  | 
|
888  | 
apply (simp add: multirel_def)  | 
|
889  | 
apply (rule trancl_type [THEN subset_trans])  | 
|
890  | 
apply (auto dest: multirel1_type [THEN subsetD])  | 
|
891  | 
done  | 
|
892  | 
||
893  | 
(* Monotonicity of multirel *)  | 
|
894  | 
lemma multirel_mono:  | 
|
895  | 
"[| A\<subseteq>B; r\<subseteq>s |] ==> multirel(A, r)\<subseteq>multirel(B,s)"  | 
|
896  | 
apply (simp add: multirel_def)  | 
|
897  | 
apply (rule trancl_mono)  | 
|
898  | 
apply (rule multirel1_mono, auto)  | 
|
899  | 
done  | 
|
900  | 
||
901  | 
(* Equivalence of multirel with the usual (closure-free) def *)  | 
|
902  | 
||
903  | 
lemma add_diff_eq: "k \<in> nat ==> 0 < k --> n #+ k #- 1 = n #+ (k #- 1)"  | 
|
904  | 
by (erule nat_induct, auto)  | 
|
905  | 
||
906  | 
lemma mdiff_union_single_conv: "[|a \<in> mset_of(J); multiset(I); multiset(J) |]  | 
|
907  | 
   ==> I +# J -# {#a#} = I +# (J-# {#a#})"
 | 
|
908  | 
apply (simp (no_asm_simp) add: multiset_equality)  | 
|
909  | 
apply (case_tac "a \<notin> mset_of (I) ")  | 
|
910  | 
apply (auto simp add: mcount_def mset_of_def multiset_def multiset_fun_iff)  | 
|
911  | 
apply (auto dest: domain_type simp add: add_diff_eq)  | 
|
912  | 
done  | 
|
913  | 
||
914  | 
lemma diff_add_commute: "[| n le m; m \<in> nat; n \<in> nat; k \<in> nat |] ==> m #- n #+ k = m #+ k #- n"  | 
|
915  | 
by (auto simp add: le_iff less_iff_succ_add)  | 
|
916  | 
||
917  | 
(* One direction *)  | 
|
918  | 
||
919  | 
lemma multirel_implies_one_step:  | 
|
920  | 
"<M,N> \<in> multirel(A, r) ==>  | 
|
921  | 
trans[A](r) -->  | 
|
922  | 
(\<exists>I J K.  | 
|
923  | 
I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) &  | 
|
924  | 
N = I +# J & M = I +# K & J \<noteq> 0 &  | 
|
925  | 
(\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r))"  | 
|
926  | 
apply (simp add: multirel_def Ball_def Bex_def)  | 
|
927  | 
apply (erule converse_trancl_induct)  | 
|
928  | 
apply (simp_all add: multirel1_iff Mult_iff_multiset)  | 
|
929  | 
(* Two subgoals remain *)  | 
|
930  | 
(* Subgoal 1 *)  | 
|
931  | 
apply clarify  | 
|
932  | 
apply (rule_tac x = M0 in exI, force)  | 
|
933  | 
(* Subgoal 2 *)  | 
|
934  | 
apply clarify  | 
|
935  | 
apply (case_tac "a \<in> mset_of (Ka) ")  | 
|
936  | 
apply (rule_tac x = I in exI, simp (no_asm_simp))  | 
|
937  | 
apply (rule_tac x = J in exI, simp (no_asm_simp))  | 
|
938  | 
apply (rule_tac x = " (Ka -# {#a#}) +# K" in exI, simp (no_asm_simp))
 | 
|
939  | 
apply (simp_all add: Un_subset_iff)  | 
|
940  | 
apply (simp (no_asm_simp) add: munion_assoc [symmetric])  | 
|
941  | 
apply (drule_tac t = "%M. M-#{#a#}" in subst_context)
 | 
|
942  | 
apply (simp add: mdiff_union_single_conv melem_diff_single, clarify)  | 
|
943  | 
apply (erule disjE, simp)  | 
|
944  | 
apply (erule disjE, simp)  | 
|
945  | 
apply (drule_tac x = a and P = "%x. x :# Ka \<longrightarrow> ?Q(x)" in spec)  | 
|
946  | 
apply clarify  | 
|
947  | 
apply (rule_tac x = xa in exI)  | 
|
948  | 
apply (simp (no_asm_simp))  | 
|
949  | 
apply (blast dest: trans_onD)  | 
|
950  | 
(* new we know that a\<notin>mset_of(Ka) *)  | 
|
951  | 
apply (subgoal_tac "a :# I")  | 
|
952  | 
apply (rule_tac x = "I-#{#a#}" in exI, simp (no_asm_simp))
 | 
|
953  | 
apply (rule_tac x = "J+#{#a#}" in exI)
 | 
|
954  | 
apply (simp (no_asm_simp) add: Un_subset_iff)  | 
|
955  | 
apply (rule_tac x = "Ka +# K" in exI)  | 
|
956  | 
apply (simp (no_asm_simp) add: Un_subset_iff)  | 
|
957  | 
apply (rule conjI)  | 
|
958  | 
apply (simp (no_asm_simp) add: multiset_equality mcount_elem [THEN succ_pred_eq_self])  | 
|
959  | 
apply (rule conjI)  | 
|
960  | 
apply (drule_tac t = "%M. M-#{#a#}" in subst_context)
 | 
|
961  | 
apply (simp add: mdiff_union_inverse2)  | 
|
962  | 
apply (simp_all (no_asm_simp) add: multiset_equality)  | 
|
963  | 
apply (rule diff_add_commute [symmetric])  | 
|
964  | 
apply (auto intro: mcount_elem)  | 
|
965  | 
apply (subgoal_tac "a \<in> mset_of (I +# Ka) ")  | 
|
966  | 
apply (drule_tac [2] sym, auto)  | 
|
967  | 
done  | 
|
968  | 
||
969  | 
lemma melem_imp_eq_diff_union [simp]: "[| a \<in> mset_of(M); multiset(M) |] ==> M -# {#a#} +# {#a#} = M"
 | 
|
970  | 
by (simp add: multiset_equality mcount_elem [THEN succ_pred_eq_self])  | 
|
971  | 
||
972  | 
lemma msize_eq_succ_imp_eq_union:  | 
|
973  | 
"[| msize(M)=$# succ(n); M \<in> Mult(A); n \<in> nat |]  | 
|
974  | 
      ==> \<exists>a N. M = N +# {#a#} & N \<in> Mult(A) & a \<in> A"
 | 
|
975  | 
apply (drule msize_eq_succ_imp_elem, auto)  | 
|
976  | 
apply (rule_tac x = a in exI)  | 
|
977  | 
apply (rule_tac x = "M -# {#a#}" in exI)
 | 
|
978  | 
apply (frule Mult_into_multiset)  | 
|
979  | 
apply (simp (no_asm_simp))  | 
|
980  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
981  | 
done  | 
|
982  | 
||
983  | 
(* The second direction *)  | 
|
984  | 
||
985  | 
lemma one_step_implies_multirel_lemma [rule_format (no_asm)]:  | 
|
986  | 
"n \<in> nat ==>  | 
|
987  | 
(\<forall>I J K.  | 
|
988  | 
I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) &  | 
|
989  | 
(msize(J) = $# n & J \<noteq>0 & (\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k, j> \<in> r))  | 
|
990  | 
--> <I +# K, I +# J> \<in> multirel(A, r))"  | 
|
991  | 
apply (simp add: Mult_iff_multiset)  | 
|
992  | 
apply (erule nat_induct, clarify)  | 
|
993  | 
apply (drule_tac M = J in msize_eq_0_iff, auto)  | 
|
994  | 
(* one subgoal remains *)  | 
|
995  | 
apply (subgoal_tac "msize (J) =$# succ (x) ")  | 
|
996  | 
prefer 2 apply simp  | 
|
997  | 
apply (frule_tac A = A in msize_eq_succ_imp_eq_union)  | 
|
998  | 
apply (simp_all add: Mult_iff_multiset, clarify)  | 
|
999  | 
apply (rename_tac "J'", simp)  | 
|
1000  | 
apply (case_tac "J' = 0")  | 
|
1001  | 
apply (simp add: multirel_def)  | 
|
1002  | 
apply (rule r_into_trancl, clarify)  | 
|
1003  | 
apply (simp add: multirel1_iff Mult_iff_multiset, force)  | 
|
1004  | 
(*Now we know J' \<noteq> 0*)  | 
|
1005  | 
apply (drule sym, rotate_tac -1, simp)  | 
|
1006  | 
apply (erule_tac V = "$# x = msize (J') " in thin_rl)  | 
|
1007  | 
apply (frule_tac M = K and P = "%x. <x,a> \<in> r" in multiset_partition)  | 
|
1008  | 
apply (erule_tac P = "\<forall>k \<in> mset_of (K) . ?P (k) " in rev_mp)  | 
|
1009  | 
apply (erule ssubst)  | 
|
1010  | 
apply (simp add: Ball_def, auto)  | 
|
| 15481 | 1011  | 
apply (subgoal_tac "< (I +# {# x \<in> K. <x, a> \<in> r#}) +# {# x \<in> K. <x, a> \<notin> r#}, (I +# {# x \<in> K. <x, a> \<in> r#}) +# J'> \<in> multirel(A, r) ")
 | 
| 15201 | 1012  | 
prefer 2  | 
1013  | 
 apply (drule_tac x = "I +# {# x \<in> K. <x, a> \<in> r#}" in spec)
 | 
|
1014  | 
apply (rotate_tac -1)  | 
|
1015  | 
apply (drule_tac x = "J'" in spec)  | 
|
1016  | 
apply (rotate_tac -1)  | 
|
1017  | 
 apply (drule_tac x = "{# x \<in> K. <x, a> \<notin> r#}" in spec, simp) apply blast
 | 
|
1018  | 
apply (simp add: munion_assoc [symmetric] multirel_def)  | 
|
1019  | 
apply (rule_tac b = "I +# {# x \<in> K. <x, a> \<in> r#} +# J'" in trancl_trans, blast)
 | 
|
1020  | 
apply (rule r_into_trancl)  | 
|
1021  | 
apply (simp add: multirel1_iff Mult_iff_multiset)  | 
|
1022  | 
apply (rule_tac x = a in exI)  | 
|
1023  | 
apply (simp (no_asm_simp))  | 
|
1024  | 
apply (rule_tac x = "I +# J'" in exI)  | 
|
1025  | 
apply (auto simp add: munion_ac Un_subset_iff)  | 
|
1026  | 
done  | 
|
1027  | 
||
1028  | 
lemma one_step_implies_multirel:  | 
|
1029  | 
"[| J \<noteq> 0; \<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r;  | 
|
1030  | 
I \<in> Mult(A); J \<in> Mult(A); K \<in> Mult(A) |]  | 
|
1031  | 
==> <I+#K, I+#J> \<in> multirel(A, r)"  | 
|
1032  | 
apply (subgoal_tac "multiset (J) ")  | 
|
1033  | 
prefer 2 apply (simp add: Mult_iff_multiset)  | 
|
1034  | 
apply (frule_tac M = J in msize_int_of_nat)  | 
|
1035  | 
apply (auto intro: one_step_implies_multirel_lemma)  | 
|
1036  | 
done  | 
|
1037  | 
||
1038  | 
(** Proving that multisets are partially ordered **)  | 
|
1039  | 
||
1040  | 
(*irreflexivity*)  | 
|
1041  | 
||
1042  | 
lemma multirel_irrefl_lemma:  | 
|
1043  | 
"Finite(A) ==> part_ord(A, r) --> (\<forall>x \<in> A. \<exists>y \<in> A. <x,y> \<in> r) -->A=0"  | 
|
1044  | 
apply (erule Finite_induct)  | 
|
1045  | 
apply (auto dest: subset_consI [THEN [2] part_ord_subset])  | 
|
1046  | 
apply (auto simp add: part_ord_def irrefl_def)  | 
|
1047  | 
apply (drule_tac x = xa in bspec)  | 
|
1048  | 
apply (drule_tac [2] a = xa and b = x in trans_onD, auto)  | 
|
1049  | 
done  | 
|
1050  | 
||
1051  | 
lemma irrefl_on_multirel:  | 
|
1052  | 
"part_ord(A, r) ==> irrefl(Mult(A), multirel(A, r))"  | 
|
1053  | 
apply (simp add: irrefl_def)  | 
|
1054  | 
apply (subgoal_tac "trans[A](r) ")  | 
|
1055  | 
prefer 2 apply (simp add: part_ord_def, clarify)  | 
|
1056  | 
apply (drule multirel_implies_one_step, clarify)  | 
|
1057  | 
apply (simp add: Mult_iff_multiset, clarify)  | 
|
1058  | 
apply (subgoal_tac "Finite (mset_of (K))")  | 
|
1059  | 
apply (frule_tac r = r in multirel_irrefl_lemma)  | 
|
1060  | 
apply (frule_tac B = "mset_of (K) " in part_ord_subset)  | 
|
1061  | 
apply simp_all  | 
|
1062  | 
apply (auto simp add: multiset_def mset_of_def)  | 
|
1063  | 
done  | 
|
1064  | 
||
1065  | 
lemma trans_on_multirel: "trans[Mult(A)](multirel(A, r))"  | 
|
1066  | 
apply (simp add: multirel_def trans_on_def)  | 
|
1067  | 
apply (blast intro: trancl_trans)  | 
|
1068  | 
done  | 
|
1069  | 
||
1070  | 
lemma multirel_trans:  | 
|
1071  | 
"[| <M, N> \<in> multirel(A, r); <N, K> \<in> multirel(A, r) |] ==> <M, K> \<in> multirel(A,r)"  | 
|
1072  | 
apply (simp add: multirel_def)  | 
|
1073  | 
apply (blast intro: trancl_trans)  | 
|
1074  | 
done  | 
|
1075  | 
||
1076  | 
lemma trans_multirel: "trans(multirel(A,r))"  | 
|
1077  | 
apply (simp add: multirel_def)  | 
|
1078  | 
apply (rule trans_trancl)  | 
|
1079  | 
done  | 
|
1080  | 
||
1081  | 
lemma part_ord_multirel: "part_ord(A,r) ==> part_ord(Mult(A), multirel(A, r))"  | 
|
1082  | 
apply (simp (no_asm) add: part_ord_def)  | 
|
1083  | 
apply (blast intro: irrefl_on_multirel trans_on_multirel)  | 
|
1084  | 
done  | 
|
1085  | 
||
1086  | 
(** Monotonicity of multiset union **)  | 
|
1087  | 
||
1088  | 
lemma munion_multirel1_mono:  | 
|
1089  | 
"[|<M,N> \<in> multirel1(A, r); K \<in> Mult(A) |] ==> <K +# M, K +# N> \<in> multirel1(A, r)"  | 
|
1090  | 
apply (frule multirel1_type [THEN subsetD])  | 
|
1091  | 
apply (auto simp add: multirel1_iff Mult_iff_multiset)  | 
|
1092  | 
apply (rule_tac x = a in exI)  | 
|
1093  | 
apply (simp (no_asm_simp))  | 
|
1094  | 
apply (rule_tac x = "K+#M0" in exI)  | 
|
1095  | 
apply (simp (no_asm_simp) add: Un_subset_iff)  | 
|
1096  | 
apply (rule_tac x = Ka in exI)  | 
|
1097  | 
apply (simp (no_asm_simp) add: munion_assoc)  | 
|
1098  | 
done  | 
|
1099  | 
||
1100  | 
lemma munion_multirel_mono2:  | 
|
1101  | 
"[| <M, N> \<in> multirel(A, r); K \<in> Mult(A) |]==><K +# M, K +# N> \<in> multirel(A, r)"  | 
|
1102  | 
apply (frule multirel_type [THEN subsetD])  | 
|
1103  | 
apply (simp (no_asm_use) add: multirel_def)  | 
|
1104  | 
apply clarify  | 
|
1105  | 
apply (drule_tac psi = "<M,N> \<in> multirel1 (A, r) ^+" in asm_rl)  | 
|
1106  | 
apply (erule rev_mp)  | 
|
1107  | 
apply (erule rev_mp)  | 
|
1108  | 
apply (erule rev_mp)  | 
|
1109  | 
apply (erule trancl_induct, clarify)  | 
|
1110  | 
apply (blast intro: munion_multirel1_mono r_into_trancl, clarify)  | 
|
| 15481 | 1111  | 
apply (subgoal_tac "y \<in> Mult(A) ")  | 
| 15201 | 1112  | 
prefer 2  | 
1113  | 
apply (blast dest: multirel_type [unfolded multirel_def, THEN subsetD])  | 
|
1114  | 
apply (subgoal_tac "<K +# y, K +# z> \<in> multirel1 (A, r) ")  | 
|
1115  | 
prefer 2 apply (blast intro: munion_multirel1_mono)  | 
|
1116  | 
apply (blast intro: r_into_trancl trancl_trans)  | 
|
1117  | 
done  | 
|
1118  | 
||
1119  | 
lemma munion_multirel_mono1:  | 
|
1120  | 
"[|<M, N> \<in> multirel(A, r); K \<in> Mult(A)|] ==> <M +# K, N +# K> \<in> multirel(A, r)"  | 
|
1121  | 
apply (frule multirel_type [THEN subsetD])  | 
|
| 15481 | 1122  | 
apply (rule_tac P = "%x. <x,?u> \<in> multirel(A, r) " in munion_commute [THEN subst])  | 
1123  | 
apply (subst munion_commute [of N])  | 
|
| 15201 | 1124  | 
apply (rule munion_multirel_mono2)  | 
1125  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
1126  | 
done  | 
|
1127  | 
||
1128  | 
lemma munion_multirel_mono:  | 
|
1129  | 
"[|<M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)|]  | 
|
1130  | 
==> <M +# N, K +# L> \<in> multirel(A, r)"  | 
|
| 15481 | 1131  | 
apply (subgoal_tac "M \<in> Mult(A) & N \<in> Mult(A) & K \<in> Mult(A) & L \<in> Mult(A) ")  | 
| 15201 | 1132  | 
prefer 2 apply (blast dest: multirel_type [THEN subsetD])  | 
1133  | 
apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2)  | 
|
1134  | 
done  | 
|
1135  | 
||
1136  | 
||
1137  | 
subsection{*Ordinal Multisets*}
 | 
|
1138  | 
||
1139  | 
(* A \<subseteq> B ==> field(Memrel(A)) \<subseteq> field(Memrel(B)) *)  | 
|
1140  | 
lemmas field_Memrel_mono = Memrel_mono [THEN field_mono, standard]  | 
|
1141  | 
||
1142  | 
(*  | 
|
1143  | 
[| Aa \<subseteq> Ba; A \<subseteq> B |] ==>  | 
|
1144  | 
multirel(field(Memrel(Aa)), Memrel(A))\<subseteq> multirel(field(Memrel(Ba)), Memrel(B))  | 
|
1145  | 
*)  | 
|
1146  | 
||
1147  | 
lemmas multirel_Memrel_mono = multirel_mono [OF field_Memrel_mono Memrel_mono]  | 
|
1148  | 
||
1149  | 
lemma omultiset_is_multiset [simp]: "omultiset(M) ==> multiset(M)"  | 
|
1150  | 
apply (simp add: omultiset_def)  | 
|
1151  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
1152  | 
done  | 
|
1153  | 
||
1154  | 
lemma munion_omultiset [simp]: "[| omultiset(M); omultiset(N) |] ==> omultiset(M +# N)"  | 
|
1155  | 
apply (simp add: omultiset_def, clarify)  | 
|
1156  | 
apply (rule_tac x = "i Un ia" in exI)  | 
|
1157  | 
apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff)  | 
|
1158  | 
apply (blast intro: field_Memrel_mono)  | 
|
1159  | 
done  | 
|
1160  | 
||
1161  | 
lemma mdiff_omultiset [simp]: "omultiset(M) ==> omultiset(M -# N)"  | 
|
1162  | 
apply (simp add: omultiset_def, clarify)  | 
|
1163  | 
apply (simp add: Mult_iff_multiset)  | 
|
1164  | 
apply (rule_tac x = i in exI)  | 
|
1165  | 
apply (simp (no_asm_simp))  | 
|
1166  | 
done  | 
|
1167  | 
||
1168  | 
(** Proving that Memrel is a partial order **)  | 
|
1169  | 
||
1170  | 
lemma irrefl_Memrel: "Ord(i) ==> irrefl(field(Memrel(i)), Memrel(i))"  | 
|
1171  | 
apply (rule irreflI, clarify)  | 
|
1172  | 
apply (subgoal_tac "Ord (x) ")  | 
|
1173  | 
prefer 2 apply (blast intro: Ord_in_Ord)  | 
|
1174  | 
apply (drule_tac i = x in ltI [THEN lt_irrefl], auto)  | 
|
1175  | 
done  | 
|
1176  | 
||
1177  | 
lemma trans_iff_trans_on: "trans(r) <-> trans[field(r)](r)"  | 
|
1178  | 
by (simp add: trans_on_def trans_def, auto)  | 
|
1179  | 
||
1180  | 
lemma part_ord_Memrel: "Ord(i) ==>part_ord(field(Memrel(i)), Memrel(i))"  | 
|
1181  | 
apply (simp add: part_ord_def)  | 
|
1182  | 
apply (simp (no_asm) add: trans_iff_trans_on [THEN iff_sym])  | 
|
1183  | 
apply (blast intro: trans_Memrel irrefl_Memrel)  | 
|
1184  | 
done  | 
|
1185  | 
||
1186  | 
(*  | 
|
1187  | 
Ord(i) ==>  | 
|
1188  | 
  part_ord(field(Memrel(i))-||>nat-{0}, multirel(field(Memrel(i)), Memrel(i)))
 | 
|
1189  | 
*)  | 
|
1190  | 
||
1191  | 
lemmas part_ord_mless = part_ord_Memrel [THEN part_ord_multirel, standard]  | 
|
1192  | 
||
1193  | 
(*irreflexivity*)  | 
|
1194  | 
||
1195  | 
lemma mless_not_refl: "~(M <# M)"  | 
|
1196  | 
apply (simp add: mless_def, clarify)  | 
|
1197  | 
apply (frule multirel_type [THEN subsetD])  | 
|
1198  | 
apply (drule part_ord_mless)  | 
|
1199  | 
apply (simp add: part_ord_def irrefl_def)  | 
|
1200  | 
done  | 
|
1201  | 
||
1202  | 
(* N<N ==> R *)  | 
|
1203  | 
lemmas mless_irrefl = mless_not_refl [THEN notE, standard, elim!]  | 
|
1204  | 
||
1205  | 
(*transitivity*)  | 
|
1206  | 
lemma mless_trans: "[| K <# M; M <# N |] ==> K <# N"  | 
|
1207  | 
apply (simp add: mless_def, clarify)  | 
|
1208  | 
apply (rule_tac x = "i Un ia" in exI)  | 
|
1209  | 
apply (blast dest: multirel_Memrel_mono [OF Un_upper1 Un_upper1, THEN subsetD]  | 
|
1210  | 
multirel_Memrel_mono [OF Un_upper2 Un_upper2, THEN subsetD]  | 
|
1211  | 
intro: multirel_trans Ord_Un)  | 
|
1212  | 
done  | 
|
1213  | 
||
1214  | 
(*asymmetry*)  | 
|
1215  | 
lemma mless_not_sym: "M <# N ==> ~ N <# M"  | 
|
1216  | 
apply clarify  | 
|
1217  | 
apply (rule mless_not_refl [THEN notE])  | 
|
1218  | 
apply (erule mless_trans, assumption)  | 
|
1219  | 
done  | 
|
1220  | 
||
1221  | 
lemma mless_asym: "[| M <# N; ~P ==> N <# M |] ==> P"  | 
|
1222  | 
by (blast dest: mless_not_sym)  | 
|
1223  | 
||
1224  | 
lemma mle_refl [simp]: "omultiset(M) ==> M <#= M"  | 
|
1225  | 
by (simp add: mle_def)  | 
|
1226  | 
||
1227  | 
(*anti-symmetry*)  | 
|
1228  | 
lemma mle_antisym:  | 
|
1229  | 
"[| M <#= N; N <#= M |] ==> M = N"  | 
|
1230  | 
apply (simp add: mle_def)  | 
|
1231  | 
apply (blast dest: mless_not_sym)  | 
|
1232  | 
done  | 
|
1233  | 
||
1234  | 
(*transitivity*)  | 
|
1235  | 
lemma mle_trans: "[| K <#= M; M <#= N |] ==> K <#= N"  | 
|
1236  | 
apply (simp add: mle_def)  | 
|
1237  | 
apply (blast intro: mless_trans)  | 
|
1238  | 
done  | 
|
1239  | 
||
1240  | 
lemma mless_le_iff: "M <# N <-> (M <#= N & M \<noteq> N)"  | 
|
1241  | 
by (simp add: mle_def, auto)  | 
|
1242  | 
||
1243  | 
(** Monotonicity of mless **)  | 
|
1244  | 
||
1245  | 
lemma munion_less_mono2: "[| M <# N; omultiset(K) |] ==> K +# M <# K +# N"  | 
|
1246  | 
apply (simp add: mless_def omultiset_def, clarify)  | 
|
1247  | 
apply (rule_tac x = "i Un ia" in exI)  | 
|
1248  | 
apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff)  | 
|
1249  | 
apply (rule munion_multirel_mono2)  | 
|
1250  | 
apply (blast intro: multirel_Memrel_mono [THEN subsetD])  | 
|
1251  | 
apply (simp add: Mult_iff_multiset)  | 
|
1252  | 
apply (blast intro: field_Memrel_mono [THEN subsetD])  | 
|
1253  | 
done  | 
|
1254  | 
||
1255  | 
lemma munion_less_mono1: "[| M <# N; omultiset(K) |] ==> M +# K <# N +# K"  | 
|
1256  | 
by (force dest: munion_less_mono2 simp add: munion_commute)  | 
|
1257  | 
||
1258  | 
lemma mless_imp_omultiset: "M <# N ==> omultiset(M) & omultiset(N)"  | 
|
1259  | 
by (auto simp add: mless_def omultiset_def dest: multirel_type [THEN subsetD])  | 
|
1260  | 
||
1261  | 
lemma munion_less_mono: "[| M <# K; N <# L |] ==> M +# N <# K +# L"  | 
|
1262  | 
apply (frule_tac M = M in mless_imp_omultiset)  | 
|
1263  | 
apply (frule_tac M = N in mless_imp_omultiset)  | 
|
1264  | 
apply (blast intro: munion_less_mono1 munion_less_mono2 mless_trans)  | 
|
1265  | 
done  | 
|
1266  | 
||
1267  | 
(* <#= *)  | 
|
1268  | 
||
1269  | 
lemma mle_imp_omultiset: "M <#= N ==> omultiset(M) & omultiset(N)"  | 
|
1270  | 
by (auto simp add: mle_def mless_imp_omultiset)  | 
|
1271  | 
||
1272  | 
lemma mle_mono: "[| M <#= K; N <#= L |] ==> M +# N <#= K +# L"  | 
|
1273  | 
apply (frule_tac M = M in mle_imp_omultiset)  | 
|
1274  | 
apply (frule_tac M = N in mle_imp_omultiset)  | 
|
1275  | 
apply (auto simp add: mle_def intro: munion_less_mono1 munion_less_mono2 munion_less_mono)  | 
|
1276  | 
done  | 
|
1277  | 
||
1278  | 
lemma omultiset_0 [iff]: "omultiset(0)"  | 
|
1279  | 
by (auto simp add: omultiset_def Mult_iff_multiset)  | 
|
1280  | 
||
1281  | 
lemma empty_leI [simp]: "omultiset(M) ==> 0 <#= M"  | 
|
1282  | 
apply (simp add: mle_def mless_def)  | 
|
| 15481 | 1283  | 
apply (subgoal_tac "\<exists>i. Ord (i) & M \<in> Mult(field(Memrel(i))) ")  | 
| 15201 | 1284  | 
prefer 2 apply (simp add: omultiset_def)  | 
1285  | 
apply (case_tac "M=0", simp_all, clarify)  | 
|
| 15481 | 1286  | 
apply (subgoal_tac "<0 +# 0, 0 +# M> \<in> multirel(field (Memrel(i)), Memrel(i))")  | 
| 15201 | 1287  | 
apply (rule_tac [2] one_step_implies_multirel)  | 
1288  | 
apply (auto simp add: Mult_iff_multiset)  | 
|
1289  | 
done  | 
|
1290  | 
||
1291  | 
lemma munion_upper1: "[| omultiset(M); omultiset(N) |] ==> M <#= M +# N"  | 
|
1292  | 
apply (subgoal_tac "M +# 0 <#= M +# N")  | 
|
1293  | 
apply (rule_tac [2] mle_mono, auto)  | 
|
1294  | 
done  | 
|
1295  | 
||
| 
12089
 
34e7693271a9
Sidi Ehmety's port of the fold_set operator and multisets to ZF.
 
paulson 
parents:  
diff
changeset
 | 
1296  | 
end  |