12 begin |
12 begin |
13 |
13 |
14 abbreviation (input) |
14 abbreviation (input) |
15 \<comment> \<open>Short cut for multiset space\<close> |
15 \<comment> \<open>Short cut for multiset space\<close> |
16 Mult :: "i=>i" where |
16 Mult :: "i=>i" where |
17 "Mult(A) == A -||> nat-{0}" |
17 "Mult(A) \<equiv> A -||> nat-{0}" |
18 |
18 |
19 definition |
19 definition |
20 (* This is the original "restrict" from ZF.thy. |
20 (* This is the original "restrict" from ZF.thy. |
21 Restricts the function f to the domain A |
21 Restricts the function f to the domain A |
22 FIXME: adapt Multiset to the new "restrict". *) |
22 FIXME: adapt Multiset to the new "restrict". *) |
23 funrestrict :: "[i,i] => i" where |
23 funrestrict :: "[i,i] => i" where |
24 "funrestrict(f,A) == \<lambda>x \<in> A. f`x" |
24 "funrestrict(f,A) \<equiv> \<lambda>x \<in> A. f`x" |
25 |
25 |
26 definition |
26 definition |
27 (* M is a multiset *) |
27 (* M is a multiset *) |
28 multiset :: "i => o" where |
28 multiset :: "i => o" where |
29 "multiset(M) == \<exists>A. M \<in> A -> nat-{0} & Finite(A)" |
29 "multiset(M) \<equiv> \<exists>A. M \<in> A -> nat-{0} & Finite(A)" |
30 |
30 |
31 definition |
31 definition |
32 mset_of :: "i=>i" where |
32 mset_of :: "i=>i" where |
33 "mset_of(M) == domain(M)" |
33 "mset_of(M) \<equiv> domain(M)" |
34 |
34 |
35 definition |
35 definition |
36 munion :: "[i, i] => i" (infixl \<open>+#\<close> 65) where |
36 munion :: "[i, i] => i" (infixl \<open>+#\<close> 65) where |
37 "M +# N == \<lambda>x \<in> mset_of(M) \<union> mset_of(N). |
37 "M +# N \<equiv> \<lambda>x \<in> mset_of(M) \<union> mset_of(N). |
38 if x \<in> mset_of(M) \<inter> mset_of(N) then (M`x) #+ (N`x) |
38 if x \<in> mset_of(M) \<inter> mset_of(N) then (M`x) #+ (N`x) |
39 else (if x \<in> mset_of(M) then M`x else N`x)" |
39 else (if x \<in> mset_of(M) then M`x else N`x)" |
40 |
40 |
41 definition |
41 definition |
42 (*convert a function to a multiset by eliminating 0*) |
42 (*convert a function to a multiset by eliminating 0*) |
43 normalize :: "i => i" where |
43 normalize :: "i => i" where |
44 "normalize(f) == |
44 "normalize(f) \<equiv> |
45 if (\<exists>A. f \<in> A -> nat & Finite(A)) then |
45 if (\<exists>A. f \<in> A -> nat & Finite(A)) then |
46 funrestrict(f, {x \<in> mset_of(f). 0 < f`x}) |
46 funrestrict(f, {x \<in> mset_of(f). 0 < f`x}) |
47 else 0" |
47 else 0" |
48 |
48 |
49 definition |
49 definition |
50 mdiff :: "[i, i] => i" (infixl \<open>-#\<close> 65) where |
50 mdiff :: "[i, i] => i" (infixl \<open>-#\<close> 65) where |
51 "M -# N == normalize(\<lambda>x \<in> mset_of(M). |
51 "M -# N \<equiv> normalize(\<lambda>x \<in> mset_of(M). |
52 if x \<in> mset_of(N) then M`x #- N`x else M`x)" |
52 if x \<in> mset_of(N) then M`x #- N`x else M`x)" |
53 |
53 |
54 definition |
54 definition |
55 (* set of elements of a multiset *) |
55 (* set of elements of a multiset *) |
56 msingle :: "i => i" (\<open>{#_#}\<close>) where |
56 msingle :: "i => i" (\<open>{#_#}\<close>) where |
57 "{#a#} == {<a, 1>}" |
57 "{#a#} \<equiv> {<a, 1>}" |
58 |
58 |
59 definition |
59 definition |
60 MCollect :: "[i, i=>o] => i" (*comprehension*) where |
60 MCollect :: "[i, i=>o] => i" (*comprehension*) where |
61 "MCollect(M, P) == funrestrict(M, {x \<in> mset_of(M). P(x)})" |
61 "MCollect(M, P) \<equiv> funrestrict(M, {x \<in> mset_of(M). P(x)})" |
62 |
62 |
63 definition |
63 definition |
64 (* Counts the number of occurrences of an element in a multiset *) |
64 (* Counts the number of occurrences of an element in a multiset *) |
65 mcount :: "[i, i] => i" where |
65 mcount :: "[i, i] => i" where |
66 "mcount(M, a) == if a \<in> mset_of(M) then M`a else 0" |
66 "mcount(M, a) \<equiv> if a \<in> mset_of(M) then M`a else 0" |
67 |
67 |
68 definition |
68 definition |
69 msize :: "i => i" where |
69 msize :: "i => i" where |
70 "msize(M) == setsum(%a. $# mcount(M,a), mset_of(M))" |
70 "msize(M) \<equiv> setsum(%a. $# mcount(M,a), mset_of(M))" |
71 |
71 |
72 abbreviation |
72 abbreviation |
73 melem :: "[i,i] => o" (\<open>(_/ :# _)\<close> [50, 51] 50) where |
73 melem :: "[i,i] => o" (\<open>(_/ :# _)\<close> [50, 51] 50) where |
74 "a :# M == a \<in> mset_of(M)" |
74 "a :# M \<equiv> a \<in> mset_of(M)" |
75 |
75 |
76 syntax |
76 syntax |
77 "_MColl" :: "[pttrn, i, o] => i" (\<open>(1{# _ \<in> _./ _#})\<close>) |
77 "_MColl" :: "[pttrn, i, o] => i" (\<open>(1{# _ \<in> _./ _#})\<close>) |
78 translations |
78 translations |
79 "{#x \<in> M. P#}" == "CONST MCollect(M, \<lambda>x. P)" |
79 "{#x \<in> M. P#}" == "CONST MCollect(M, \<lambda>x. P)" |
82 |
82 |
83 definition |
83 definition |
84 (* multirel1 has to be a set (not a predicate) so that we can form |
84 (* multirel1 has to be a set (not a predicate) so that we can form |
85 its transitive closure and reason about wf(.) and acc(.) *) |
85 its transitive closure and reason about wf(.) and acc(.) *) |
86 multirel1 :: "[i,i]=>i" where |
86 multirel1 :: "[i,i]=>i" where |
87 "multirel1(A, r) == |
87 "multirel1(A, r) \<equiv> |
88 {<M, N> \<in> Mult(A)*Mult(A). |
88 {<M, N> \<in> Mult(A)*Mult(A). |
89 \<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A). |
89 \<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A). |
90 N=M0 +# {#a#} & M=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}" |
90 N=M0 +# {#a#} & M=M0 +# K & (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}" |
91 |
91 |
92 definition |
92 definition |
93 multirel :: "[i, i] => i" where |
93 multirel :: "[i, i] => i" where |
94 "multirel(A, r) == multirel1(A, r)^+" |
94 "multirel(A, r) \<equiv> multirel1(A, r)^+" |
95 |
95 |
96 (* ordinal multiset orderings *) |
96 (* ordinal multiset orderings *) |
97 |
97 |
98 definition |
98 definition |
99 omultiset :: "i => o" where |
99 omultiset :: "i => o" where |
100 "omultiset(M) == \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))" |
100 "omultiset(M) \<equiv> \<exists>i. Ord(i) & M \<in> Mult(field(Memrel(i)))" |
101 |
101 |
102 definition |
102 definition |
103 mless :: "[i, i] => o" (infixl \<open><#\<close> 50) where |
103 mless :: "[i, i] => o" (infixl \<open><#\<close> 50) where |
104 "M <# N == \<exists>i. Ord(i) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))" |
104 "M <# N \<equiv> \<exists>i. Ord(i) & <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))" |
105 |
105 |
106 definition |
106 definition |
107 mle :: "[i, i] => o" (infixl \<open><#=\<close> 50) where |
107 mle :: "[i, i] => o" (infixl \<open><#=\<close> 50) where |
108 "M <#= N == (omultiset(M) & M = N) | M <# N" |
108 "M <#= N \<equiv> (omultiset(M) & M = N) | M <# N" |
109 |
109 |
110 |
110 |
111 subsection\<open>Properties of the original "restrict" from ZF.thy\<close> |
111 subsection\<open>Properties of the original "restrict" from ZF.thy\<close> |
112 |
112 |
113 lemma funrestrict_subset: "[| f \<in> Pi(C,B); A\<subseteq>C |] ==> funrestrict(f,A) \<subseteq> f" |
113 lemma funrestrict_subset: "\<lbrakk>f \<in> Pi(C,B); A\<subseteq>C\<rbrakk> \<Longrightarrow> funrestrict(f,A) \<subseteq> f" |
114 by (auto simp add: funrestrict_def lam_def intro: apply_Pair) |
114 by (auto simp add: funrestrict_def lam_def intro: apply_Pair) |
115 |
115 |
116 lemma funrestrict_type: |
116 lemma funrestrict_type: |
117 "[| !!x. x \<in> A ==> f`x \<in> B(x) |] ==> funrestrict(f,A) \<in> Pi(A,B)" |
117 "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> f`x \<in> B(x)\<rbrakk> \<Longrightarrow> funrestrict(f,A) \<in> Pi(A,B)" |
118 by (simp add: funrestrict_def lam_type) |
118 by (simp add: funrestrict_def lam_type) |
119 |
119 |
120 lemma funrestrict_type2: "[| f \<in> Pi(C,B); A\<subseteq>C |] ==> funrestrict(f,A) \<in> Pi(A,B)" |
120 lemma funrestrict_type2: "\<lbrakk>f \<in> Pi(C,B); A\<subseteq>C\<rbrakk> \<Longrightarrow> funrestrict(f,A) \<in> Pi(A,B)" |
121 by (blast intro: apply_type funrestrict_type) |
121 by (blast intro: apply_type funrestrict_type) |
122 |
122 |
123 lemma funrestrict [simp]: "a \<in> A ==> funrestrict(f,A) ` a = f`a" |
123 lemma funrestrict [simp]: "a \<in> A \<Longrightarrow> funrestrict(f,A) ` a = f`a" |
124 by (simp add: funrestrict_def) |
124 by (simp add: funrestrict_def) |
125 |
125 |
126 lemma funrestrict_empty [simp]: "funrestrict(f,0) = 0" |
126 lemma funrestrict_empty [simp]: "funrestrict(f,0) = 0" |
127 by (simp add: funrestrict_def) |
127 by (simp add: funrestrict_def) |
128 |
128 |
129 lemma domain_funrestrict [simp]: "domain(funrestrict(f,C)) = C" |
129 lemma domain_funrestrict [simp]: "domain(funrestrict(f,C)) = C" |
130 by (auto simp add: funrestrict_def lam_def) |
130 by (auto simp add: funrestrict_def lam_def) |
131 |
131 |
132 lemma fun_cons_funrestrict_eq: |
132 lemma fun_cons_funrestrict_eq: |
133 "f \<in> cons(a, b) -> B ==> f = cons(<a, f ` a>, funrestrict(f, b))" |
133 "f \<in> cons(a, b) -> B \<Longrightarrow> f = cons(<a, f ` a>, funrestrict(f, b))" |
134 apply (rule equalityI) |
134 apply (rule equalityI) |
135 prefer 2 apply (blast intro: apply_Pair funrestrict_subset [THEN subsetD]) |
135 prefer 2 apply (blast intro: apply_Pair funrestrict_subset [THEN subsetD]) |
136 apply (auto dest!: Pi_memberD simp add: funrestrict_def lam_def) |
136 apply (auto dest!: Pi_memberD simp add: funrestrict_def lam_def) |
137 done |
137 done |
138 |
138 |
301 apply (force intro!: lam_type) |
301 apply (force intro!: lam_type) |
302 done |
302 done |
303 |
303 |
304 (** Count of elements **) |
304 (** Count of elements **) |
305 |
305 |
306 lemma mcount_type [simp,TC]: "multiset(M) ==> mcount(M, a) \<in> nat" |
306 lemma mcount_type [simp,TC]: "multiset(M) \<Longrightarrow> mcount(M, a) \<in> nat" |
307 by (auto simp add: multiset_def mcount_def mset_of_def multiset_fun_iff) |
307 by (auto simp add: multiset_def mcount_def mset_of_def multiset_fun_iff) |
308 |
308 |
309 lemma mcount_0 [simp]: "mcount(0, a) = 0" |
309 lemma mcount_0 [simp]: "mcount(0, a) = 0" |
310 by (simp add: mcount_def) |
310 by (simp add: mcount_def) |
311 |
311 |
312 lemma mcount_single [simp]: "mcount({#b#}, a) = (if a=b then 1 else 0)" |
312 lemma mcount_single [simp]: "mcount({#b#}, a) = (if a=b then 1 else 0)" |
313 by (simp add: mcount_def mset_of_def msingle_def) |
313 by (simp add: mcount_def mset_of_def msingle_def) |
314 |
314 |
315 lemma mcount_union [simp]: "[| multiset(M); multiset(N) |] |
315 lemma mcount_union [simp]: "\<lbrakk>multiset(M); multiset(N)\<rbrakk> |
316 ==> mcount(M +# N, a) = mcount(M, a) #+ mcount (N, a)" |
316 \<Longrightarrow> mcount(M +# N, a) = mcount(M, a) #+ mcount (N, a)" |
317 apply (auto simp add: multiset_def multiset_fun_iff mcount_def munion_def mset_of_def) |
317 apply (auto simp add: multiset_def multiset_fun_iff mcount_def munion_def mset_of_def) |
318 done |
318 done |
319 |
319 |
320 lemma mcount_diff [simp]: |
320 lemma mcount_diff [simp]: |
321 "multiset(M) ==> mcount(M -# N, a) = mcount(M, a) #- mcount(N, a)" |
321 "multiset(M) \<Longrightarrow> mcount(M -# N, a) = mcount(M, a) #- mcount(N, a)" |
322 apply (simp add: multiset_def) |
322 apply (simp add: multiset_def) |
323 apply (auto dest!: not_lt_imp_le |
323 apply (auto dest!: not_lt_imp_le |
324 simp add: mdiff_def multiset_fun_iff mcount_def normalize_def mset_of_def) |
324 simp add: mdiff_def multiset_fun_iff mcount_def normalize_def mset_of_def) |
325 apply (force intro!: lam_type) |
325 apply (force intro!: lam_type) |
326 apply (force intro!: lam_type) |
326 apply (force intro!: lam_type) |
327 done |
327 done |
328 |
328 |
329 lemma mcount_elem: "[| multiset(M); a \<in> mset_of(M) |] ==> 0 < mcount(M, a)" |
329 lemma mcount_elem: "\<lbrakk>multiset(M); a \<in> mset_of(M)\<rbrakk> \<Longrightarrow> 0 < mcount(M, a)" |
330 apply (simp add: multiset_def, clarify) |
330 apply (simp add: multiset_def, clarify) |
331 apply (simp add: mcount_def mset_of_def) |
331 apply (simp add: mcount_def mset_of_def) |
332 apply (simp add: multiset_fun_iff) |
332 apply (simp add: multiset_fun_iff) |
333 done |
333 done |
334 |
334 |
374 apply (rule not_zneg_int_of [THEN bexE]) |
374 apply (rule not_zneg_int_of [THEN bexE]) |
375 apply (auto simp del: int_of_0 simp add: int_of_add [symmetric] int_of_0 [symmetric]) |
375 apply (auto simp del: int_of_0 simp add: int_of_add [symmetric] int_of_0 [symmetric]) |
376 done |
376 done |
377 |
377 |
378 lemma setsum_mcount_Int: |
378 lemma setsum_mcount_Int: |
379 "Finite(A) ==> setsum(%a. $# mcount(N, a), A \<inter> mset_of(N)) |
379 "Finite(A) \<Longrightarrow> setsum(%a. $# mcount(N, a), A \<inter> mset_of(N)) |
380 = setsum(%a. $# mcount(N, a), A)" |
380 = setsum(%a. $# mcount(N, a), A)" |
381 apply (induct rule: Finite_induct) |
381 apply (induct rule: Finite_induct) |
382 apply auto |
382 apply auto |
383 apply (subgoal_tac "Finite (B \<inter> mset_of (N))") |
383 apply (subgoal_tac "Finite (B \<inter> mset_of (N))") |
384 prefer 2 apply (blast intro: subset_Finite) |
384 prefer 2 apply (blast intro: subset_Finite) |
385 apply (auto simp add: mcount_def Int_cons_left) |
385 apply (auto simp add: mcount_def Int_cons_left) |
386 done |
386 done |
387 |
387 |
388 lemma msize_union [simp]: |
388 lemma msize_union [simp]: |
389 "[| multiset(M); multiset(N) |] ==> msize(M +# N) = msize(M) $+ msize(N)" |
389 "\<lbrakk>multiset(M); multiset(N)\<rbrakk> \<Longrightarrow> msize(M +# N) = msize(M) $+ msize(N)" |
390 apply (simp add: msize_def setsum_Un setsum_addf int_of_add setsum_mcount_Int) |
390 apply (simp add: msize_def setsum_Un setsum_addf int_of_add setsum_mcount_Int) |
391 apply (subst Int_commute) |
391 apply (subst Int_commute) |
392 apply (simp add: setsum_mcount_Int) |
392 apply (simp add: setsum_mcount_Int) |
393 done |
393 done |
394 |
394 |
395 lemma msize_eq_succ_imp_elem: "[|msize(M)= $# succ(n); n \<in> nat|] ==> \<exists>a. a \<in> mset_of(M)" |
395 lemma msize_eq_succ_imp_elem: "\<lbrakk>msize(M)= $# succ(n); n \<in> nat\<rbrakk> \<Longrightarrow> \<exists>a. a \<in> mset_of(M)" |
396 apply (unfold msize_def) |
396 apply (unfold msize_def) |
397 apply (blast dest: setsum_succD) |
397 apply (blast dest: setsum_succD) |
398 done |
398 done |
399 |
399 |
400 (** Equality of multisets **) |
400 (** Equality of multisets **) |
401 |
401 |
402 lemma equality_lemma: |
402 lemma equality_lemma: |
403 "[| multiset(M); multiset(N); \<forall>a. mcount(M, a)=mcount(N, a) |] |
403 "\<lbrakk>multiset(M); multiset(N); \<forall>a. mcount(M, a)=mcount(N, a)\<rbrakk> |
404 ==> mset_of(M)=mset_of(N)" |
404 \<Longrightarrow> mset_of(M)=mset_of(N)" |
405 apply (simp add: multiset_def) |
405 apply (simp add: multiset_def) |
406 apply (rule sym, rule equalityI) |
406 apply (rule sym, rule equalityI) |
407 apply (auto simp add: multiset_fun_iff mcount_def mset_of_def) |
407 apply (auto simp add: multiset_fun_iff mcount_def mset_of_def) |
408 apply (drule_tac [!] x=x in spec) |
408 apply (drule_tac [!] x=x in spec) |
409 apply (case_tac [2] "x \<in> Aa", case_tac "x \<in> A", auto) |
409 apply (case_tac [2] "x \<in> Aa", case_tac "x \<in> A", auto) |
410 done |
410 done |
411 |
411 |
412 lemma multiset_equality: |
412 lemma multiset_equality: |
413 "[| multiset(M); multiset(N) |]==> M=N\<longleftrightarrow>(\<forall>a. mcount(M, a)=mcount(N, a))" |
413 "\<lbrakk>multiset(M); multiset(N)\<rbrakk>\<Longrightarrow> M=N\<longleftrightarrow>(\<forall>a. mcount(M, a)=mcount(N, a))" |
414 apply auto |
414 apply auto |
415 apply (subgoal_tac "mset_of (M) = mset_of (N) ") |
415 apply (subgoal_tac "mset_of (M) = mset_of (N) ") |
416 prefer 2 apply (blast intro: equality_lemma) |
416 prefer 2 apply (blast intro: equality_lemma) |
417 apply (simp add: multiset_def mset_of_def) |
417 apply (simp add: multiset_def mset_of_def) |
418 apply (auto simp add: multiset_fun_iff) |
418 apply (auto simp add: multiset_fun_iff) |
490 apply (rule int_of_diff, auto) |
490 apply (rule int_of_diff, auto) |
491 done |
491 done |
492 |
492 |
493 lemma setsum_decr2: |
493 lemma setsum_decr2: |
494 "Finite(A) |
494 "Finite(A) |
495 ==> \<forall>M. multiset(M) \<longrightarrow> (\<forall>a \<in> mset_of(M). |
495 \<Longrightarrow> \<forall>M. multiset(M) \<longrightarrow> (\<forall>a \<in> mset_of(M). |
496 setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) = |
496 setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) = |
497 (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a |
497 (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a |
498 else setsum(%x. $# mcount(M, x), A)))" |
498 else setsum(%x. $# mcount(M, x), A)))" |
499 apply (simp add: multiset_def) |
499 apply (simp add: multiset_def) |
500 apply (erule Finite_induct) |
500 apply (erule Finite_induct) |
501 apply (auto simp add: multiset_fun_iff mcount_def mset_of_def) |
501 apply (auto simp add: multiset_fun_iff mcount_def mset_of_def) |
502 done |
502 done |
503 |
503 |
504 lemma setsum_decr3: "[| Finite(A); multiset(M); a \<in> mset_of(M) |] |
504 lemma setsum_decr3: "\<lbrakk>Finite(A); multiset(M); a \<in> mset_of(M)\<rbrakk> |
505 ==> setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) = |
505 \<Longrightarrow> setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) = |
506 (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a |
506 (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a |
507 else setsum(%x. $# mcount(M, x), A))" |
507 else setsum(%x. $# mcount(M, x), A))" |
508 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) ") |
508 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) ") |
509 apply (rule_tac [2] setsum_Diff [symmetric]) |
509 apply (rule_tac [2] setsum_Diff [symmetric]) |
510 apply (rule sym, rule ssubst, blast) |
510 apply (rule sym, rule ssubst, blast) |
511 apply (rule sym, drule setsum_decr2, auto) |
511 apply (rule sym, drule setsum_decr2, auto) |
512 apply (simp add: mcount_def mset_of_def) |
512 apply (simp add: mcount_def mset_of_def) |
513 done |
513 done |
514 |
514 |
515 lemma nat_le_1_cases: "n \<in> nat ==> n \<le> 1 \<longleftrightarrow> (n=0 | n=1)" |
515 lemma nat_le_1_cases: "n \<in> nat \<Longrightarrow> n \<le> 1 \<longleftrightarrow> (n=0 | n=1)" |
516 by (auto elim: natE) |
516 by (auto elim: natE) |
517 |
517 |
518 lemma succ_pred_eq_self: "[| 0<n; n \<in> nat |] ==> succ(n #- 1) = n" |
518 lemma succ_pred_eq_self: "\<lbrakk>0<n; n \<in> nat\<rbrakk> \<Longrightarrow> succ(n #- 1) = n" |
519 apply (subgoal_tac "1 \<le> n") |
519 apply (subgoal_tac "1 \<le> n") |
520 apply (drule add_diff_inverse2, auto) |
520 apply (drule add_diff_inverse2, auto) |
521 done |
521 done |
522 |
522 |
523 text\<open>Specialized for use in the proof below.\<close> |
523 text\<open>Specialized for use in the proof below.\<close> |
598 apply (subgoal_tac "{x \<in> A - {a} . 0 < funrestrict (M, A - {x}) ` x} = A - {a}") |
598 apply (subgoal_tac "{x \<in> A - {a} . 0 < funrestrict (M, A - {x}) ` x} = A - {a}") |
599 apply (auto intro!: setsum_cong simp add: zdiff_eq_iff zadd_commute multiset_def multiset_fun_iff mset_of_def) |
599 apply (auto intro!: setsum_cong simp add: zdiff_eq_iff zadd_commute multiset_def multiset_fun_iff mset_of_def) |
600 done |
600 done |
601 |
601 |
602 lemma multiset_induct2: |
602 lemma multiset_induct2: |
603 "[| multiset(M); P(0); |
603 "\<lbrakk>multiset(M); P(0); |
604 (!!M a. [| multiset(M); a\<notin>mset_of(M); P(M) |] ==> P(cons(<a, 1>, M))); |
604 (\<And>M a. \<lbrakk>multiset(M); a\<notin>mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(cons(<a, 1>, M))); |
605 (!!M b. [| multiset(M); b \<in> mset_of(M); P(M) |] ==> P(M(b:= M`b #+ 1))) |] |
605 (\<And>M b. \<lbrakk>multiset(M); b \<in> mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(M(b:= M`b #+ 1)))\<rbrakk> |
606 ==> P(M)" |
606 \<Longrightarrow> P(M)" |
607 apply (subgoal_tac "\<exists>n \<in> nat. setsum (\<lambda>x. $# mcount (M, x), {x \<in> mset_of (M) . 0 < M ` x}) = $# n") |
607 apply (subgoal_tac "\<exists>n \<in> nat. setsum (\<lambda>x. $# mcount (M, x), {x \<in> mset_of (M) . 0 < M ` x}) = $# n") |
608 apply (rule_tac [2] not_zneg_int_of) |
608 apply (rule_tac [2] not_zneg_int_of) |
609 apply (simp_all (no_asm_simp) add: znegative_iff_zless_0 not_zless_iff_zle) |
609 apply (simp_all (no_asm_simp) add: znegative_iff_zless_0 not_zless_iff_zle) |
610 apply (rule_tac [2] g_zpos_imp_setsum_zpos) |
610 apply (rule_tac [2] g_zpos_imp_setsum_zpos) |
611 prefer 2 apply (blast intro: multiset_set_of_Finite Collect_subset [THEN subset_Finite]) |
611 prefer 2 apply (blast intro: multiset_set_of_Finite Collect_subset [THEN subset_Finite]) |
612 prefer 2 apply (simp add: multiset_def multiset_fun_iff, clarify) |
612 prefer 2 apply (simp add: multiset_def multiset_fun_iff, clarify) |
613 apply (rule multiset_induct_aux [rule_format], auto) |
613 apply (rule multiset_induct_aux [rule_format], auto) |
614 done |
614 done |
615 |
615 |
616 lemma munion_single_case1: |
616 lemma munion_single_case1: |
617 "[| multiset(M); a \<notin>mset_of(M) |] ==> M +# {#a#} = cons(<a, 1>, M)" |
617 "\<lbrakk>multiset(M); a \<notin>mset_of(M)\<rbrakk> \<Longrightarrow> M +# {#a#} = cons(<a, 1>, M)" |
618 apply (simp add: multiset_def msingle_def) |
618 apply (simp add: multiset_def msingle_def) |
619 apply (auto simp add: munion_def) |
619 apply (auto simp add: munion_def) |
620 apply (unfold mset_of_def, simp) |
620 apply (unfold mset_of_def, simp) |
621 apply (rule fun_extension, rule lam_type, simp_all) |
621 apply (rule fun_extension, rule lam_type, simp_all) |
622 apply (auto simp add: multiset_fun_iff fun_extend_apply) |
622 apply (auto simp add: multiset_fun_iff fun_extend_apply) |
623 apply (drule_tac c = a and b = 1 in fun_extend3) |
623 apply (drule_tac c = a and b = 1 in fun_extend3) |
624 apply (auto simp add: cons_eq Un_commute [of _ "{a}"]) |
624 apply (auto simp add: cons_eq Un_commute [of _ "{a}"]) |
625 done |
625 done |
626 |
626 |
627 lemma munion_single_case2: |
627 lemma munion_single_case2: |
628 "[| multiset(M); a \<in> mset_of(M) |] ==> M +# {#a#} = M(a:=M`a #+ 1)" |
628 "\<lbrakk>multiset(M); a \<in> mset_of(M)\<rbrakk> \<Longrightarrow> M +# {#a#} = M(a:=M`a #+ 1)" |
629 apply (simp add: multiset_def) |
629 apply (simp add: multiset_def) |
630 apply (auto simp add: munion_def multiset_fun_iff msingle_def) |
630 apply (auto simp add: munion_def multiset_fun_iff msingle_def) |
631 apply (unfold mset_of_def, simp) |
631 apply (unfold mset_of_def, simp) |
632 apply (subgoal_tac "A \<union> {a} = A") |
632 apply (subgoal_tac "A \<union> {a} = A") |
633 apply (rule fun_extension) |
633 apply (rule fun_extension) |
649 done |
649 done |
650 |
650 |
651 (** MCollect **) |
651 (** MCollect **) |
652 |
652 |
653 lemma MCollect_multiset [simp]: |
653 lemma MCollect_multiset [simp]: |
654 "multiset(M) ==> multiset({# x \<in> M. P(x)#})" |
654 "multiset(M) \<Longrightarrow> multiset({# x \<in> M. P(x)#})" |
655 apply (simp add: MCollect_def multiset_def mset_of_def, clarify) |
655 apply (simp add: MCollect_def multiset_def mset_of_def, clarify) |
656 apply (rule_tac x = "{x \<in> A. P (x) }" in exI) |
656 apply (rule_tac x = "{x \<in> A. P (x) }" in exI) |
657 apply (auto dest: CollectD1 [THEN [2] apply_type] |
657 apply (auto dest: CollectD1 [THEN [2] apply_type] |
658 intro: Collect_subset [THEN subset_Finite] funrestrict_type) |
658 intro: Collect_subset [THEN subset_Finite] funrestrict_type) |
659 done |
659 done |
660 |
660 |
661 lemma mset_of_MCollect [simp]: |
661 lemma mset_of_MCollect [simp]: |
662 "multiset(M) ==> mset_of({# x \<in> M. P(x) #}) \<subseteq> mset_of(M)" |
662 "multiset(M) \<Longrightarrow> mset_of({# x \<in> M. P(x) #}) \<subseteq> mset_of(M)" |
663 by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def) |
663 by (auto simp add: mset_of_def MCollect_def multiset_def funrestrict_def) |
664 |
664 |
665 lemma MCollect_mem_iff [iff]: |
665 lemma MCollect_mem_iff [iff]: |
666 "x \<in> mset_of({#x \<in> M. P(x)#}) \<longleftrightarrow> x \<in> mset_of(M) & P(x)" |
666 "x \<in> mset_of({#x \<in> M. P(x)#}) \<longleftrightarrow> x \<in> mset_of(M) & P(x)" |
667 by (simp add: MCollect_def mset_of_def) |
667 by (simp add: MCollect_def mset_of_def) |
668 |
668 |
669 lemma mcount_MCollect [simp]: |
669 lemma mcount_MCollect [simp]: |
670 "mcount({# x \<in> M. P(x) #}, a) = (if P(a) then mcount(M,a) else 0)" |
670 "mcount({# x \<in> M. P(x) #}, a) = (if P(a) then mcount(M,a) else 0)" |
671 by (simp add: mcount_def MCollect_def mset_of_def) |
671 by (simp add: mcount_def MCollect_def mset_of_def) |
672 |
672 |
673 lemma multiset_partition: "multiset(M) ==> M = {# x \<in> M. P(x) #} +# {# x \<in> M. ~ P(x) #}" |
673 lemma multiset_partition: "multiset(M) \<Longrightarrow> M = {# x \<in> M. P(x) #} +# {# x \<in> M. \<not> P(x) #}" |
674 by (simp add: multiset_equality) |
674 by (simp add: multiset_equality) |
675 |
675 |
676 lemma natify_elem_is_self [simp]: |
676 lemma natify_elem_is_self [simp]: |
677 "[| multiset(M); a \<in> mset_of(M) |] ==> natify(M`a) = M`a" |
677 "\<lbrakk>multiset(M); a \<in> mset_of(M)\<rbrakk> \<Longrightarrow> natify(M`a) = M`a" |
678 by (auto simp add: multiset_def mset_of_def multiset_fun_iff) |
678 by (auto simp add: multiset_def mset_of_def multiset_fun_iff) |
679 |
679 |
680 (* and more algebraic laws on multisets *) |
680 (* and more algebraic laws on multisets *) |
681 |
681 |
682 lemma munion_eq_conv_diff: "[| multiset(M); multiset(N) |] |
682 lemma munion_eq_conv_diff: "\<lbrakk>multiset(M); multiset(N)\<rbrakk> |
683 ==> (M +# {#a#} = N +# {#b#}) \<longleftrightarrow> (M = N & a = b | |
683 \<Longrightarrow> (M +# {#a#} = N +# {#b#}) \<longleftrightarrow> (M = N & a = b | |
684 M = N -# {#a#} +# {#b#} & N = M -# {#b#} +# {#a#})" |
684 M = N -# {#a#} +# {#b#} & N = M -# {#b#} +# {#a#})" |
685 apply (simp del: mcount_single add: multiset_equality) |
685 apply (simp del: mcount_single add: multiset_equality) |
686 apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE) |
686 apply (rule iffI, erule_tac [2] disjE, erule_tac [3] conjE) |
687 apply (case_tac "a=b", auto) |
687 apply (case_tac "a=b", auto) |
688 apply (drule_tac x = a in spec) |
688 apply (drule_tac x = a in spec) |
733 by (auto simp add: multirel1_def Mult_iff_multiset Bex_def) |
733 by (auto simp add: multirel1_def Mult_iff_multiset Bex_def) |
734 |
734 |
735 |
735 |
736 text\<open>Monotonicity of \<^term>\<open>multirel1\<close>\<close> |
736 text\<open>Monotonicity of \<^term>\<open>multirel1\<close>\<close> |
737 |
737 |
738 lemma multirel1_mono1: "A\<subseteq>B ==> multirel1(A, r)\<subseteq>multirel1(B, r)" |
738 lemma multirel1_mono1: "A\<subseteq>B \<Longrightarrow> multirel1(A, r)\<subseteq>multirel1(B, r)" |
739 apply (auto simp add: multirel1_def) |
739 apply (auto simp add: multirel1_def) |
740 apply (auto simp add: Un_subset_iff Mult_iff_multiset) |
740 apply (auto simp add: Un_subset_iff Mult_iff_multiset) |
741 apply (rule_tac x = a in bexI) |
741 apply (rule_tac x = a in bexI) |
742 apply (rule_tac x = M0 in bexI, simp) |
742 apply (rule_tac x = M0 in bexI, simp) |
743 apply (rule_tac x = K in bexI) |
743 apply (rule_tac x = K in bexI) |
744 apply (auto simp add: Mult_iff_multiset) |
744 apply (auto simp add: Mult_iff_multiset) |
745 done |
745 done |
746 |
746 |
747 lemma multirel1_mono2: "r\<subseteq>s ==> multirel1(A,r)\<subseteq>multirel1(A, s)" |
747 lemma multirel1_mono2: "r\<subseteq>s \<Longrightarrow> multirel1(A,r)\<subseteq>multirel1(A, s)" |
748 apply (simp add: multirel1_def, auto) |
748 apply (simp add: multirel1_def, auto) |
749 apply (rule_tac x = a in bexI) |
749 apply (rule_tac x = a in bexI) |
750 apply (rule_tac x = M0 in bexI) |
750 apply (rule_tac x = M0 in bexI) |
751 apply (simp_all add: Mult_iff_multiset) |
751 apply (simp_all add: Mult_iff_multiset) |
752 apply (rule_tac x = K in bexI) |
752 apply (rule_tac x = K in bexI) |
753 apply (simp_all add: Mult_iff_multiset, auto) |
753 apply (simp_all add: Mult_iff_multiset, auto) |
754 done |
754 done |
755 |
755 |
756 lemma multirel1_mono: |
756 lemma multirel1_mono: |
757 "[| A\<subseteq>B; r\<subseteq>s |] ==> multirel1(A, r) \<subseteq> multirel1(B, s)" |
757 "\<lbrakk>A\<subseteq>B; r\<subseteq>s\<rbrakk> \<Longrightarrow> multirel1(A, r) \<subseteq> multirel1(B, s)" |
758 apply (rule subset_trans) |
758 apply (rule subset_trans) |
759 apply (rule multirel1_mono1) |
759 apply (rule multirel1_mono1) |
760 apply (rule_tac [2] multirel1_mono2, auto) |
760 apply (rule_tac [2] multirel1_mono2, auto) |
761 done |
761 done |
762 |
762 |
763 subsection\<open>Toward the proof of well-foundedness of multirel1\<close> |
763 subsection\<open>Toward the proof of well-foundedness of multirel1\<close> |
764 |
764 |
765 lemma not_less_0 [iff]: "<M,0> \<notin> multirel1(A, r)" |
765 lemma not_less_0 [iff]: "<M,0> \<notin> multirel1(A, r)" |
766 by (auto simp add: multirel1_def Mult_iff_multiset) |
766 by (auto simp add: multirel1_def Mult_iff_multiset) |
767 |
767 |
768 lemma less_munion: "[| <N, M0 +# {#a#}> \<in> multirel1(A, r); M0 \<in> Mult(A) |] ==> |
768 lemma less_munion: "\<lbrakk><N, M0 +# {#a#}> \<in> multirel1(A, r); M0 \<in> Mult(A)\<rbrakk> \<Longrightarrow> |
769 (\<exists>M. <M, M0> \<in> multirel1(A, r) & N = M +# {#a#}) | |
769 (\<exists>M. <M, M0> \<in> multirel1(A, r) & N = M +# {#a#}) | |
770 (\<exists>K. K \<in> Mult(A) & (\<forall>b \<in> mset_of(K). <b, a> \<in> r) & N = M0 +# K)" |
770 (\<exists>K. K \<in> Mult(A) & (\<forall>b \<in> mset_of(K). <b, a> \<in> r) & N = M0 +# K)" |
771 apply (frule multirel1_type [THEN subsetD]) |
771 apply (frule multirel1_type [THEN subsetD]) |
772 apply (simp add: multirel1_iff) |
772 apply (simp add: multirel1_iff) |
773 apply (auto simp add: munion_eq_conv_exist) |
773 apply (auto simp add: munion_eq_conv_exist) |
774 apply (rule_tac x="Ka +# K" in exI, auto, simp add: Mult_iff_multiset) |
774 apply (rule_tac x="Ka +# K" in exI, auto, simp add: Mult_iff_multiset) |
775 apply (simp (no_asm_simp) add: munion_left_cancel munion_assoc) |
775 apply (simp (no_asm_simp) add: munion_left_cancel munion_assoc) |
776 apply (auto simp add: munion_commute) |
776 apply (auto simp add: munion_commute) |
777 done |
777 done |
778 |
778 |
779 lemma multirel1_base: "[| M \<in> Mult(A); a \<in> A |] ==> <M, M +# {#a#}> \<in> multirel1(A, r)" |
779 lemma multirel1_base: "\<lbrakk>M \<in> Mult(A); a \<in> A\<rbrakk> \<Longrightarrow> <M, M +# {#a#}> \<in> multirel1(A, r)" |
780 apply (auto simp add: multirel1_iff) |
780 apply (auto simp add: multirel1_iff) |
781 apply (simp add: Mult_iff_multiset) |
781 apply (simp add: Mult_iff_multiset) |
782 apply (rule_tac x = a in exI, clarify) |
782 apply (rule_tac x = a in exI, clarify) |
783 apply (rule_tac x = M in exI, simp) |
783 apply (rule_tac x = M in exI, simp) |
784 apply (rule_tac x = 0 in exI, auto) |
784 apply (rule_tac x = 0 in exI, auto) |
785 done |
785 done |
786 |
786 |
787 lemma acc_0: "acc(0)=0" |
787 lemma acc_0: "acc(0)=0" |
788 by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD]) |
788 by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD]) |
789 |
789 |
790 lemma lemma1: "[| \<forall>b \<in> A. <b,a> \<in> r \<longrightarrow> |
790 lemma lemma1: "\<lbrakk>\<forall>b \<in> A. <b,a> \<in> r \<longrightarrow> |
791 (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r))); |
791 (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r))); |
792 M0 \<in> acc(multirel1(A, r)); a \<in> A; |
792 M0 \<in> acc(multirel1(A, r)); a \<in> A; |
793 \<forall>M. <M,M0> \<in> multirel1(A, r) \<longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r)) |] |
793 \<forall>M. <M,M0> \<in> multirel1(A, r) \<longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r))\<rbrakk> |
794 ==> M0 +# {#a#} \<in> acc(multirel1(A, r))" |
794 \<Longrightarrow> M0 +# {#a#} \<in> acc(multirel1(A, r))" |
795 apply (subgoal_tac "M0 \<in> Mult(A) ") |
795 apply (subgoal_tac "M0 \<in> Mult(A) ") |
796 prefer 2 |
796 prefer 2 |
797 apply (erule acc.cases) |
797 apply (erule acc.cases) |
798 apply (erule fieldE) |
798 apply (erule fieldE) |
799 apply (auto dest: multirel1_type [THEN subsetD]) |
799 apply (auto dest: multirel1_type [THEN subsetD]) |
887 apply (auto dest: multirel1_type [THEN subsetD]) |
887 apply (auto dest: multirel1_type [THEN subsetD]) |
888 done |
888 done |
889 |
889 |
890 (* Monotonicity of multirel *) |
890 (* Monotonicity of multirel *) |
891 lemma multirel_mono: |
891 lemma multirel_mono: |
892 "[| A\<subseteq>B; r\<subseteq>s |] ==> multirel(A, r)\<subseteq>multirel(B,s)" |
892 "\<lbrakk>A\<subseteq>B; r\<subseteq>s\<rbrakk> \<Longrightarrow> multirel(A, r)\<subseteq>multirel(B,s)" |
893 apply (simp add: multirel_def) |
893 apply (simp add: multirel_def) |
894 apply (rule trancl_mono) |
894 apply (rule trancl_mono) |
895 apply (rule multirel1_mono, auto) |
895 apply (rule multirel1_mono, auto) |
896 done |
896 done |
897 |
897 |
898 (* Equivalence of multirel with the usual (closure-free) definition *) |
898 (* Equivalence of multirel with the usual (closure-free) definition *) |
899 |
899 |
900 lemma add_diff_eq: "k \<in> nat ==> 0 < k \<longrightarrow> n #+ k #- 1 = n #+ (k #- 1)" |
900 lemma add_diff_eq: "k \<in> nat \<Longrightarrow> 0 < k \<longrightarrow> n #+ k #- 1 = n #+ (k #- 1)" |
901 by (erule nat_induct, auto) |
901 by (erule nat_induct, auto) |
902 |
902 |
903 lemma mdiff_union_single_conv: "[|a \<in> mset_of(J); multiset(I); multiset(J) |] |
903 lemma mdiff_union_single_conv: "\<lbrakk>a \<in> mset_of(J); multiset(I); multiset(J)\<rbrakk> |
904 ==> I +# J -# {#a#} = I +# (J-# {#a#})" |
904 \<Longrightarrow> I +# J -# {#a#} = I +# (J-# {#a#})" |
905 apply (simp (no_asm_simp) add: multiset_equality) |
905 apply (simp (no_asm_simp) add: multiset_equality) |
906 apply (case_tac "a \<notin> mset_of (I) ") |
906 apply (case_tac "a \<notin> mset_of (I) ") |
907 apply (auto simp add: mcount_def mset_of_def multiset_def multiset_fun_iff) |
907 apply (auto simp add: mcount_def mset_of_def multiset_def multiset_fun_iff) |
908 apply (auto dest: domain_type simp add: add_diff_eq) |
908 apply (auto dest: domain_type simp add: add_diff_eq) |
909 done |
909 done |
910 |
910 |
911 lemma diff_add_commute: "[| n \<le> m; m \<in> nat; n \<in> nat; k \<in> nat |] ==> m #- n #+ k = m #+ k #- n" |
911 lemma diff_add_commute: "\<lbrakk>n \<le> m; m \<in> nat; n \<in> nat; k \<in> nat\<rbrakk> \<Longrightarrow> m #- n #+ k = m #+ k #- n" |
912 by (auto simp add: le_iff less_iff_succ_add) |
912 by (auto simp add: le_iff less_iff_succ_add) |
913 |
913 |
914 (* One direction *) |
914 (* One direction *) |
915 |
915 |
916 lemma multirel_implies_one_step: |
916 lemma multirel_implies_one_step: |
917 "<M,N> \<in> multirel(A, r) ==> |
917 "<M,N> \<in> multirel(A, r) \<Longrightarrow> |
918 trans[A](r) \<longrightarrow> |
918 trans[A](r) \<longrightarrow> |
919 (\<exists>I J K. |
919 (\<exists>I J K. |
920 I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) & |
920 I \<in> Mult(A) & J \<in> Mult(A) & K \<in> Mult(A) & |
921 N = I +# J & M = I +# K & J \<noteq> 0 & |
921 N = I +# J & M = I +# K & J \<noteq> 0 & |
922 (\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r))" |
922 (\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r))" |
1113 prefer 2 apply (blast intro: munion_multirel1_mono) |
1113 prefer 2 apply (blast intro: munion_multirel1_mono) |
1114 apply (blast intro: r_into_trancl trancl_trans) |
1114 apply (blast intro: r_into_trancl trancl_trans) |
1115 done |
1115 done |
1116 |
1116 |
1117 lemma munion_multirel_mono1: |
1117 lemma munion_multirel_mono1: |
1118 "[|<M, N> \<in> multirel(A, r); K \<in> Mult(A)|] ==> <M +# K, N +# K> \<in> multirel(A, r)" |
1118 "\<lbrakk><M, N> \<in> multirel(A, r); K \<in> Mult(A)\<rbrakk> \<Longrightarrow> <M +# K, N +# K> \<in> multirel(A, r)" |
1119 apply (frule multirel_type [THEN subsetD]) |
1119 apply (frule multirel_type [THEN subsetD]) |
1120 apply (rule_tac P = "%x. <x,u> \<in> multirel(A, r)" for u in munion_commute [THEN subst]) |
1120 apply (rule_tac P = "%x. <x,u> \<in> multirel(A, r)" for u in munion_commute [THEN subst]) |
1121 apply (subst munion_commute [of N]) |
1121 apply (subst munion_commute [of N]) |
1122 apply (rule munion_multirel_mono2) |
1122 apply (rule munion_multirel_mono2) |
1123 apply (auto simp add: Mult_iff_multiset) |
1123 apply (auto simp add: Mult_iff_multiset) |
1124 done |
1124 done |
1125 |
1125 |
1126 lemma munion_multirel_mono: |
1126 lemma munion_multirel_mono: |
1127 "[|<M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)|] |
1127 "\<lbrakk><M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)\<rbrakk> |
1128 ==> <M +# N, K +# L> \<in> multirel(A, r)" |
1128 \<Longrightarrow> <M +# N, K +# L> \<in> multirel(A, r)" |
1129 apply (subgoal_tac "M \<in> Mult(A) & N \<in> Mult(A) & K \<in> Mult(A) & L \<in> Mult(A) ") |
1129 apply (subgoal_tac "M \<in> Mult(A) & N \<in> Mult(A) & K \<in> Mult(A) & L \<in> Mult(A) ") |
1130 prefer 2 apply (blast dest: multirel_type [THEN subsetD]) |
1130 prefer 2 apply (blast dest: multirel_type [THEN subsetD]) |
1131 apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2) |
1131 apply (blast intro: munion_multirel_mono1 multirel_trans munion_multirel_mono2) |
1132 done |
1132 done |
1133 |
1133 |
1134 |
1134 |
1135 subsection\<open>Ordinal Multisets\<close> |
1135 subsection\<open>Ordinal Multisets\<close> |
1136 |
1136 |
1137 (* A \<subseteq> B ==> field(Memrel(A)) \<subseteq> field(Memrel(B)) *) |
1137 (* A \<subseteq> B \<Longrightarrow> field(Memrel(A)) \<subseteq> field(Memrel(B)) *) |
1138 lemmas field_Memrel_mono = Memrel_mono [THEN field_mono] |
1138 lemmas field_Memrel_mono = Memrel_mono [THEN field_mono] |
1139 |
1139 |
1140 (* |
1140 (* |
1141 [| Aa \<subseteq> Ba; A \<subseteq> B |] ==> |
1141 \<lbrakk>Aa \<subseteq> Ba; A \<subseteq> B\<rbrakk> \<Longrightarrow> |
1142 multirel(field(Memrel(Aa)), Memrel(A))\<subseteq> multirel(field(Memrel(Ba)), Memrel(B)) |
1142 multirel(field(Memrel(Aa)), Memrel(A))\<subseteq> multirel(field(Memrel(Ba)), Memrel(B)) |
1143 *) |
1143 *) |
1144 |
1144 |
1145 lemmas multirel_Memrel_mono = multirel_mono [OF field_Memrel_mono Memrel_mono] |
1145 lemmas multirel_Memrel_mono = multirel_mono [OF field_Memrel_mono Memrel_mono] |
1146 |
1146 |
1147 lemma omultiset_is_multiset [simp]: "omultiset(M) ==> multiset(M)" |
1147 lemma omultiset_is_multiset [simp]: "omultiset(M) \<Longrightarrow> multiset(M)" |
1148 apply (simp add: omultiset_def) |
1148 apply (simp add: omultiset_def) |
1149 apply (auto simp add: Mult_iff_multiset) |
1149 apply (auto simp add: Mult_iff_multiset) |
1150 done |
1150 done |
1151 |
1151 |
1152 lemma munion_omultiset [simp]: "[| omultiset(M); omultiset(N) |] ==> omultiset(M +# N)" |
1152 lemma munion_omultiset [simp]: "\<lbrakk>omultiset(M); omultiset(N)\<rbrakk> \<Longrightarrow> omultiset(M +# N)" |
1153 apply (simp add: omultiset_def, clarify) |
1153 apply (simp add: omultiset_def, clarify) |
1154 apply (rule_tac x = "i \<union> ia" in exI) |
1154 apply (rule_tac x = "i \<union> ia" in exI) |
1155 apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff) |
1155 apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff) |
1156 apply (blast intro: field_Memrel_mono) |
1156 apply (blast intro: field_Memrel_mono) |
1157 done |
1157 done |
1158 |
1158 |
1159 lemma mdiff_omultiset [simp]: "omultiset(M) ==> omultiset(M -# N)" |
1159 lemma mdiff_omultiset [simp]: "omultiset(M) \<Longrightarrow> omultiset(M -# N)" |
1160 apply (simp add: omultiset_def, clarify) |
1160 apply (simp add: omultiset_def, clarify) |
1161 apply (simp add: Mult_iff_multiset) |
1161 apply (simp add: Mult_iff_multiset) |
1162 apply (rule_tac x = i in exI) |
1162 apply (rule_tac x = i in exI) |
1163 apply (simp (no_asm_simp)) |
1163 apply (simp (no_asm_simp)) |
1164 done |
1164 done |
1165 |
1165 |
1166 (** Proving that Memrel is a partial order **) |
1166 (** Proving that Memrel is a partial order **) |
1167 |
1167 |
1168 lemma irrefl_Memrel: "Ord(i) ==> irrefl(field(Memrel(i)), Memrel(i))" |
1168 lemma irrefl_Memrel: "Ord(i) \<Longrightarrow> irrefl(field(Memrel(i)), Memrel(i))" |
1169 apply (rule irreflI, clarify) |
1169 apply (rule irreflI, clarify) |
1170 apply (subgoal_tac "Ord (x) ") |
1170 apply (subgoal_tac "Ord (x) ") |
1171 prefer 2 apply (blast intro: Ord_in_Ord) |
1171 prefer 2 apply (blast intro: Ord_in_Ord) |
1172 apply (drule_tac i = x in ltI [THEN lt_irrefl], auto) |
1172 apply (drule_tac i = x in ltI [THEN lt_irrefl], auto) |
1173 done |
1173 done |
1174 |
1174 |
1175 lemma trans_iff_trans_on: "trans(r) \<longleftrightarrow> trans[field(r)](r)" |
1175 lemma trans_iff_trans_on: "trans(r) \<longleftrightarrow> trans[field(r)](r)" |
1176 by (simp add: trans_on_def trans_def, auto) |
1176 by (simp add: trans_on_def trans_def, auto) |
1177 |
1177 |
1178 lemma part_ord_Memrel: "Ord(i) ==>part_ord(field(Memrel(i)), Memrel(i))" |
1178 lemma part_ord_Memrel: "Ord(i) \<Longrightarrow>part_ord(field(Memrel(i)), Memrel(i))" |
1179 apply (simp add: part_ord_def) |
1179 apply (simp add: part_ord_def) |
1180 apply (simp (no_asm) add: trans_iff_trans_on [THEN iff_sym]) |
1180 apply (simp (no_asm) add: trans_iff_trans_on [THEN iff_sym]) |
1181 apply (blast intro: trans_Memrel irrefl_Memrel) |
1181 apply (blast intro: trans_Memrel irrefl_Memrel) |
1182 done |
1182 done |
1183 |
1183 |
1184 (* |
1184 (* |
1185 Ord(i) ==> |
1185 Ord(i) \<Longrightarrow> |
1186 part_ord(field(Memrel(i))-||>nat-{0}, multirel(field(Memrel(i)), Memrel(i))) |
1186 part_ord(field(Memrel(i))-||>nat-{0}, multirel(field(Memrel(i)), Memrel(i))) |
1187 *) |
1187 *) |
1188 |
1188 |
1189 lemmas part_ord_mless = part_ord_Memrel [THEN part_ord_multirel] |
1189 lemmas part_ord_mless = part_ord_Memrel [THEN part_ord_multirel] |
1190 |
1190 |
1191 (*irreflexivity*) |
1191 (*irreflexivity*) |
1192 |
1192 |
1193 lemma mless_not_refl: "~(M <# M)" |
1193 lemma mless_not_refl: "\<not>(M <# M)" |
1194 apply (simp add: mless_def, clarify) |
1194 apply (simp add: mless_def, clarify) |
1195 apply (frule multirel_type [THEN subsetD]) |
1195 apply (frule multirel_type [THEN subsetD]) |
1196 apply (drule part_ord_mless) |
1196 apply (drule part_ord_mless) |
1197 apply (simp add: part_ord_def irrefl_def) |
1197 apply (simp add: part_ord_def irrefl_def) |
1198 done |
1198 done |
1199 |
1199 |
1200 (* N<N ==> R *) |
1200 (* N<N \<Longrightarrow> R *) |
1201 lemmas mless_irrefl = mless_not_refl [THEN notE, elim!] |
1201 lemmas mless_irrefl = mless_not_refl [THEN notE, elim!] |
1202 |
1202 |
1203 (*transitivity*) |
1203 (*transitivity*) |
1204 lemma mless_trans: "[| K <# M; M <# N |] ==> K <# N" |
1204 lemma mless_trans: "\<lbrakk>K <# M; M <# N\<rbrakk> \<Longrightarrow> K <# N" |
1205 apply (simp add: mless_def, clarify) |
1205 apply (simp add: mless_def, clarify) |
1206 apply (rule_tac x = "i \<union> ia" in exI) |
1206 apply (rule_tac x = "i \<union> ia" in exI) |
1207 apply (blast dest: multirel_Memrel_mono [OF Un_upper1 Un_upper1, THEN subsetD] |
1207 apply (blast dest: multirel_Memrel_mono [OF Un_upper1 Un_upper1, THEN subsetD] |
1208 multirel_Memrel_mono [OF Un_upper2 Un_upper2, THEN subsetD] |
1208 multirel_Memrel_mono [OF Un_upper2 Un_upper2, THEN subsetD] |
1209 intro: multirel_trans Ord_Un) |
1209 intro: multirel_trans Ord_Un) |
1210 done |
1210 done |
1211 |
1211 |
1212 (*asymmetry*) |
1212 (*asymmetry*) |
1213 lemma mless_not_sym: "M <# N ==> ~ N <# M" |
1213 lemma mless_not_sym: "M <# N \<Longrightarrow> \<not> N <# M" |
1214 apply clarify |
1214 apply clarify |
1215 apply (rule mless_not_refl [THEN notE]) |
1215 apply (rule mless_not_refl [THEN notE]) |
1216 apply (erule mless_trans, assumption) |
1216 apply (erule mless_trans, assumption) |
1217 done |
1217 done |
1218 |
1218 |
1219 lemma mless_asym: "[| M <# N; ~P ==> N <# M |] ==> P" |
1219 lemma mless_asym: "\<lbrakk>M <# N; \<not>P \<Longrightarrow> N <# M\<rbrakk> \<Longrightarrow> P" |
1220 by (blast dest: mless_not_sym) |
1220 by (blast dest: mless_not_sym) |
1221 |
1221 |
1222 lemma mle_refl [simp]: "omultiset(M) ==> M <#= M" |
1222 lemma mle_refl [simp]: "omultiset(M) \<Longrightarrow> M <#= M" |
1223 by (simp add: mle_def) |
1223 by (simp add: mle_def) |
1224 |
1224 |
1225 (*anti-symmetry*) |
1225 (*anti-symmetry*) |
1226 lemma mle_antisym: |
1226 lemma mle_antisym: |
1227 "[| M <#= N; N <#= M |] ==> M = N" |
1227 "\<lbrakk>M <#= N; N <#= M\<rbrakk> \<Longrightarrow> M = N" |
1228 apply (simp add: mle_def) |
1228 apply (simp add: mle_def) |
1229 apply (blast dest: mless_not_sym) |
1229 apply (blast dest: mless_not_sym) |
1230 done |
1230 done |
1231 |
1231 |
1232 (*transitivity*) |
1232 (*transitivity*) |
1233 lemma mle_trans: "[| K <#= M; M <#= N |] ==> K <#= N" |
1233 lemma mle_trans: "\<lbrakk>K <#= M; M <#= N\<rbrakk> \<Longrightarrow> K <#= N" |
1234 apply (simp add: mle_def) |
1234 apply (simp add: mle_def) |
1235 apply (blast intro: mless_trans) |
1235 apply (blast intro: mless_trans) |
1236 done |
1236 done |
1237 |
1237 |
1238 lemma mless_le_iff: "M <# N \<longleftrightarrow> (M <#= N & M \<noteq> N)" |
1238 lemma mless_le_iff: "M <# N \<longleftrightarrow> (M <#= N & M \<noteq> N)" |
1239 by (simp add: mle_def, auto) |
1239 by (simp add: mle_def, auto) |
1240 |
1240 |
1241 (** Monotonicity of mless **) |
1241 (** Monotonicity of mless **) |
1242 |
1242 |
1243 lemma munion_less_mono2: "[| M <# N; omultiset(K) |] ==> K +# M <# K +# N" |
1243 lemma munion_less_mono2: "\<lbrakk>M <# N; omultiset(K)\<rbrakk> \<Longrightarrow> K +# M <# K +# N" |
1244 apply (simp add: mless_def omultiset_def, clarify) |
1244 apply (simp add: mless_def omultiset_def, clarify) |
1245 apply (rule_tac x = "i \<union> ia" in exI) |
1245 apply (rule_tac x = "i \<union> ia" in exI) |
1246 apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff) |
1246 apply (simp add: Mult_iff_multiset Ord_Un Un_subset_iff) |
1247 apply (rule munion_multirel_mono2) |
1247 apply (rule munion_multirel_mono2) |
1248 apply (blast intro: multirel_Memrel_mono [THEN subsetD]) |
1248 apply (blast intro: multirel_Memrel_mono [THEN subsetD]) |
1249 apply (simp add: Mult_iff_multiset) |
1249 apply (simp add: Mult_iff_multiset) |
1250 apply (blast intro: field_Memrel_mono [THEN subsetD]) |
1250 apply (blast intro: field_Memrel_mono [THEN subsetD]) |
1251 done |
1251 done |
1252 |
1252 |
1253 lemma munion_less_mono1: "[| M <# N; omultiset(K) |] ==> M +# K <# N +# K" |
1253 lemma munion_less_mono1: "\<lbrakk>M <# N; omultiset(K)\<rbrakk> \<Longrightarrow> M +# K <# N +# K" |
1254 by (force dest: munion_less_mono2 simp add: munion_commute) |
1254 by (force dest: munion_less_mono2 simp add: munion_commute) |
1255 |
1255 |
1256 lemma mless_imp_omultiset: "M <# N ==> omultiset(M) & omultiset(N)" |
1256 lemma mless_imp_omultiset: "M <# N \<Longrightarrow> omultiset(M) & omultiset(N)" |
1257 by (auto simp add: mless_def omultiset_def dest: multirel_type [THEN subsetD]) |
1257 by (auto simp add: mless_def omultiset_def dest: multirel_type [THEN subsetD]) |
1258 |
1258 |
1259 lemma munion_less_mono: "[| M <# K; N <# L |] ==> M +# N <# K +# L" |
1259 lemma munion_less_mono: "\<lbrakk>M <# K; N <# L\<rbrakk> \<Longrightarrow> M +# N <# K +# L" |
1260 apply (frule_tac M = M in mless_imp_omultiset) |
1260 apply (frule_tac M = M in mless_imp_omultiset) |
1261 apply (frule_tac M = N in mless_imp_omultiset) |
1261 apply (frule_tac M = N in mless_imp_omultiset) |
1262 apply (blast intro: munion_less_mono1 munion_less_mono2 mless_trans) |
1262 apply (blast intro: munion_less_mono1 munion_less_mono2 mless_trans) |
1263 done |
1263 done |
1264 |
1264 |
1265 (* <#= *) |
1265 (* <#= *) |
1266 |
1266 |
1267 lemma mle_imp_omultiset: "M <#= N ==> omultiset(M) & omultiset(N)" |
1267 lemma mle_imp_omultiset: "M <#= N \<Longrightarrow> omultiset(M) & omultiset(N)" |
1268 by (auto simp add: mle_def mless_imp_omultiset) |
1268 by (auto simp add: mle_def mless_imp_omultiset) |
1269 |
1269 |
1270 lemma mle_mono: "[| M <#= K; N <#= L |] ==> M +# N <#= K +# L" |
1270 lemma mle_mono: "\<lbrakk>M <#= K; N <#= L\<rbrakk> \<Longrightarrow> M +# N <#= K +# L" |
1271 apply (frule_tac M = M in mle_imp_omultiset) |
1271 apply (frule_tac M = M in mle_imp_omultiset) |
1272 apply (frule_tac M = N in mle_imp_omultiset) |
1272 apply (frule_tac M = N in mle_imp_omultiset) |
1273 apply (auto simp add: mle_def intro: munion_less_mono1 munion_less_mono2 munion_less_mono) |
1273 apply (auto simp add: mle_def intro: munion_less_mono1 munion_less_mono2 munion_less_mono) |
1274 done |
1274 done |
1275 |
1275 |
1276 lemma omultiset_0 [iff]: "omultiset(0)" |
1276 lemma omultiset_0 [iff]: "omultiset(0)" |
1277 by (auto simp add: omultiset_def Mult_iff_multiset) |
1277 by (auto simp add: omultiset_def Mult_iff_multiset) |
1278 |
1278 |
1279 lemma empty_leI [simp]: "omultiset(M) ==> 0 <#= M" |
1279 lemma empty_leI [simp]: "omultiset(M) \<Longrightarrow> 0 <#= M" |
1280 apply (simp add: mle_def mless_def) |
1280 apply (simp add: mle_def mless_def) |
1281 apply (subgoal_tac "\<exists>i. Ord (i) & M \<in> Mult(field(Memrel(i))) ") |
1281 apply (subgoal_tac "\<exists>i. Ord (i) & M \<in> Mult(field(Memrel(i))) ") |
1282 prefer 2 apply (simp add: omultiset_def) |
1282 prefer 2 apply (simp add: omultiset_def) |
1283 apply (case_tac "M=0", simp_all, clarify) |
1283 apply (case_tac "M=0", simp_all, clarify) |
1284 apply (subgoal_tac "<0 +# 0, 0 +# M> \<in> multirel(field (Memrel(i)), Memrel(i))") |
1284 apply (subgoal_tac "<0 +# 0, 0 +# M> \<in> multirel(field (Memrel(i)), Memrel(i))") |
1285 apply (rule_tac [2] one_step_implies_multirel) |
1285 apply (rule_tac [2] one_step_implies_multirel) |
1286 apply (auto simp add: Mult_iff_multiset) |
1286 apply (auto simp add: Mult_iff_multiset) |
1287 done |
1287 done |
1288 |
1288 |
1289 lemma munion_upper1: "[| omultiset(M); omultiset(N) |] ==> M <#= M +# N" |
1289 lemma munion_upper1: "\<lbrakk>omultiset(M); omultiset(N)\<rbrakk> \<Longrightarrow> M <#= M +# N" |
1290 apply (subgoal_tac "M +# 0 <#= M +# N") |
1290 apply (subgoal_tac "M +# 0 <#= M +# N") |
1291 apply (rule_tac [2] mle_mono, auto) |
1291 apply (rule_tac [2] mle_mono, auto) |
1292 done |
1292 done |
1293 |
1293 |
1294 end |
1294 end |