| author | wenzelm | 
| Tue, 18 Jul 2023 12:55:43 +0200 | |
| changeset 78393 | a2d22d519bf2 | 
| parent 78099 | 4d9349989d94 | 
| child 79575 | b21d8401f0ca | 
| permissions | -rw-r--r-- | 
| 10249 | 1  | 
(* Title: HOL/Library/Multiset.thy  | 
| 15072 | 2  | 
Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
3  | 
Author: Andrei Popescu, TU Muenchen  | 
| 59813 | 4  | 
Author: Jasmin Blanchette, Inria, LORIA, MPII  | 
5  | 
Author: Dmitriy Traytel, TU Muenchen  | 
|
6  | 
Author: Mathias Fleury, MPII  | 
|
| 74803 | 7  | 
Author: Martin Desharnais, MPI-INF Saarbruecken  | 
| 10249 | 8  | 
*)  | 
9  | 
||
| 65048 | 10  | 
section \<open>(Finite) Multisets\<close>  | 
| 10249 | 11  | 
|
| 15131 | 12  | 
theory Multiset  | 
| 
74865
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
13  | 
imports Cancellation  | 
| 15131 | 14  | 
begin  | 
| 10249 | 15  | 
|
| 60500 | 16  | 
subsection \<open>The type of multisets\<close>  | 
| 10249 | 17  | 
|
| 73270 | 18  | 
typedef 'a multiset = \<open>{f :: 'a \<Rightarrow> nat. finite {x. f x > 0}}\<close>
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
19  | 
morphisms count Abs_multiset  | 
| 10249 | 20  | 
proof  | 
| 73270 | 21  | 
  show \<open>(\<lambda>x. 0::nat) \<in> {f. finite {x. f x > 0}}\<close>
 | 
22  | 
by simp  | 
|
| 10249 | 23  | 
qed  | 
24  | 
||
| 
47429
 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 
bulwahn 
parents: 
47308 
diff
changeset
 | 
25  | 
setup_lifting type_definition_multiset  | 
| 19086 | 26  | 
|
| 73270 | 27  | 
lemma count_Abs_multiset:  | 
28  | 
  \<open>count (Abs_multiset f) = f\<close> if \<open>finite {x. f x > 0}\<close>
 | 
|
29  | 
by (rule Abs_multiset_inverse) (simp add: that)  | 
|
30  | 
||
| 60606 | 31  | 
lemma multiset_eq_iff: "M = N \<longleftrightarrow> (\<forall>a. count M a = count N a)"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
32  | 
by (simp only: count_inject [symmetric] fun_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
33  | 
|
| 60606 | 34  | 
lemma multiset_eqI: "(\<And>x. count A x = count B x) \<Longrightarrow> A = B"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
35  | 
using multiset_eq_iff by auto  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
36  | 
|
| 69593 | 37  | 
text \<open>Preservation of the representing set \<^term>\<open>multiset\<close>.\<close>  | 
| 60606 | 38  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
39  | 
lemma diff_preserves_multiset:  | 
| 73270 | 40  | 
  \<open>finite {x. 0 < M x - N x}\<close> if \<open>finite {x. 0 < M x}\<close> for M N :: \<open>'a \<Rightarrow> nat\<close>
 | 
41  | 
using that by (rule rev_finite_subset) auto  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
42  | 
|
| 
41069
 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 
haftmann 
parents: 
40968 
diff
changeset
 | 
43  | 
lemma filter_preserves_multiset:  | 
| 73270 | 44  | 
  \<open>finite {x. 0 < (if P x then M x else 0)}\<close> if \<open>finite {x. 0 < M x}\<close> for M N :: \<open>'a \<Rightarrow> nat\<close>
 | 
45  | 
using that by (rule rev_finite_subset) auto  | 
|
46  | 
||
47  | 
lemmas in_multiset = diff_preserves_multiset filter_preserves_multiset  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
48  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
49  | 
|
| 60500 | 50  | 
subsection \<open>Representing multisets\<close>  | 
51  | 
||
52  | 
text \<open>Multiset enumeration\<close>  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
53  | 
|
| 48008 | 54  | 
instantiation multiset :: (type) cancel_comm_monoid_add  | 
| 
25571
 
c9e39eafc7a0
instantiation target rather than legacy instance
 
haftmann 
parents: 
25507 
diff
changeset
 | 
55  | 
begin  | 
| 
 
c9e39eafc7a0
instantiation target rather than legacy instance
 
haftmann 
parents: 
25507 
diff
changeset
 | 
56  | 
|
| 73393 | 57  | 
lift_definition zero_multiset :: \<open>'a multiset\<close>  | 
58  | 
is \<open>\<lambda>a. 0\<close>  | 
|
| 73270 | 59  | 
by simp  | 
| 
25571
 
c9e39eafc7a0
instantiation target rather than legacy instance
 
haftmann 
parents: 
25507 
diff
changeset
 | 
60  | 
|
| 73393 | 61  | 
abbreviation empty_mset :: \<open>'a multiset\<close> (\<open>{#}\<close>)
 | 
62  | 
where \<open>empty_mset \<equiv> 0\<close>  | 
|
63  | 
||
64  | 
lift_definition plus_multiset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close>  | 
|
65  | 
is \<open>\<lambda>M N a. M a + N a\<close>  | 
|
| 73270 | 66  | 
by simp  | 
| 
25571
 
c9e39eafc7a0
instantiation target rather than legacy instance
 
haftmann 
parents: 
25507 
diff
changeset
 | 
67  | 
|
| 73393 | 68  | 
lift_definition minus_multiset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close>  | 
69  | 
is \<open>\<lambda>M N a. M a - N a\<close>  | 
|
| 73270 | 70  | 
by (rule diff_preserves_multiset)  | 
| 
59815
 
cce82e360c2f
explicit commutative additive inverse operation;
 
haftmann 
parents: 
59813 
diff
changeset
 | 
71  | 
|
| 48008 | 72  | 
instance  | 
| 73270 | 73  | 
by (standard; transfer) (simp_all add: fun_eq_iff)  | 
| 
25571
 
c9e39eafc7a0
instantiation target rather than legacy instance
 
haftmann 
parents: 
25507 
diff
changeset
 | 
74  | 
|
| 
 
c9e39eafc7a0
instantiation target rather than legacy instance
 
haftmann 
parents: 
25507 
diff
changeset
 | 
75  | 
end  | 
| 10249 | 76  | 
|
| 63195 | 77  | 
context  | 
78  | 
begin  | 
|
79  | 
||
80  | 
qualified definition is_empty :: "'a multiset \<Rightarrow> bool" where  | 
|
81  | 
  [code_abbrev]: "is_empty A \<longleftrightarrow> A = {#}"
 | 
|
82  | 
||
83  | 
end  | 
|
84  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
85  | 
lemma add_mset_in_multiset:  | 
| 73270 | 86  | 
  \<open>finite {x. 0 < (if x = a then Suc (M x) else M x)}\<close>
 | 
87  | 
  if \<open>finite {x. 0 < M x}\<close>
 | 
|
88  | 
using that by (simp add: flip: insert_Collect)  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
89  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
90  | 
lift_definition add_mset :: "'a \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" is  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
91  | 
"\<lambda>a M b. if b = a then Suc (M b) else M b"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
92  | 
by (rule add_mset_in_multiset)  | 
| 15869 | 93  | 
|
| 26145 | 94  | 
syntax  | 
| 60606 | 95  | 
  "_multiset" :: "args \<Rightarrow> 'a multiset"    ("{#(_)#}")
 | 
| 25507 | 96  | 
translations  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
97  | 
  "{#x, xs#}" == "CONST add_mset x {#xs#}"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
98  | 
  "{#x#}" == "CONST add_mset x {#}"
 | 
| 25507 | 99  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
100  | 
lemma count_empty [simp]: "count {#} a = 0"
 | 
| 
47429
 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 
bulwahn 
parents: 
47308 
diff
changeset
 | 
101  | 
by (simp add: zero_multiset.rep_eq)  | 
| 10249 | 102  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
103  | 
lemma count_add_mset [simp]:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
104  | 
"count (add_mset b A) a = (if b = a then Suc (count A a) else count A a)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
105  | 
by (simp add: add_mset.rep_eq)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
106  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
107  | 
lemma count_single: "count {#b#} a = (if b = a then 1 else 0)"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
108  | 
by simp  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
109  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
110  | 
lemma  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
111  | 
  add_mset_not_empty [simp]: \<open>add_mset a A \<noteq> {#}\<close> and
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
112  | 
  empty_not_add_mset [simp]: "{#} \<noteq> add_mset a A"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
113  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
114  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
115  | 
lemma add_mset_add_mset_same_iff [simp]:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
116  | 
"add_mset a A = add_mset a B \<longleftrightarrow> A = B"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
117  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
118  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
119  | 
lemma add_mset_commute:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
120  | 
"add_mset x (add_mset y M) = add_mset y (add_mset x M)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
121  | 
by (auto simp: multiset_eq_iff)  | 
| 29901 | 122  | 
|
| 10249 | 123  | 
|
| 60500 | 124  | 
subsection \<open>Basic operations\<close>  | 
125  | 
||
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
126  | 
subsubsection \<open>Conversion to set and membership\<close>  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
127  | 
|
| 73393 | 128  | 
definition set_mset :: \<open>'a multiset \<Rightarrow> 'a set\<close>  | 
129  | 
  where \<open>set_mset M = {x. count M x > 0}\<close>
 | 
|
130  | 
||
131  | 
abbreviation member_mset :: \<open>'a \<Rightarrow> 'a multiset \<Rightarrow> bool\<close>  | 
|
132  | 
where \<open>member_mset a M \<equiv> a \<in> set_mset M\<close>  | 
|
| 
62537
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
133  | 
|
| 
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
134  | 
notation  | 
| 73393 | 135  | 
member_mset (\<open>'(\<in>#')\<close>) and  | 
| 73394 | 136  | 
member_mset (\<open>(_/ \<in># _)\<close> [50, 51] 50)  | 
| 
62537
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
137  | 
|
| 
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
138  | 
notation (ASCII)  | 
| 73393 | 139  | 
member_mset (\<open>'(:#')\<close>) and  | 
| 73394 | 140  | 
member_mset (\<open>(_/ :# _)\<close> [50, 51] 50)  | 
| 73393 | 141  | 
|
142  | 
abbreviation not_member_mset :: \<open>'a \<Rightarrow> 'a multiset \<Rightarrow> bool\<close>  | 
|
143  | 
where \<open>not_member_mset a M \<equiv> a \<notin> set_mset M\<close>  | 
|
| 
62537
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
144  | 
|
| 
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
145  | 
notation  | 
| 73393 | 146  | 
not_member_mset (\<open>'(\<notin>#')\<close>) and  | 
| 73394 | 147  | 
not_member_mset (\<open>(_/ \<notin># _)\<close> [50, 51] 50)  | 
| 
62537
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
148  | 
|
| 
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
149  | 
notation (ASCII)  | 
| 73393 | 150  | 
not_member_mset (\<open>'(~:#')\<close>) and  | 
| 73394 | 151  | 
not_member_mset (\<open>(_/ ~:# _)\<close> [50, 51] 50)  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
152  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
153  | 
context  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
154  | 
begin  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
155  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
156  | 
qualified abbreviation Ball :: "'a multiset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
157  | 
where "Ball M \<equiv> Set.Ball (set_mset M)"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
158  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
159  | 
qualified abbreviation Bex :: "'a multiset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
160  | 
where "Bex M \<equiv> Set.Bex (set_mset M)"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
161  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
162  | 
end  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
163  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
164  | 
syntax  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
165  | 
  "_MBall"       :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<forall>_\<in>#_./ _)" [0, 0, 10] 10)
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
166  | 
  "_MBex"        :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<exists>_\<in>#_./ _)" [0, 0, 10] 10)
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
167  | 
|
| 
62537
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
168  | 
syntax (ASCII)  | 
| 
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
169  | 
  "_MBall"       :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<forall>_:#_./ _)" [0, 0, 10] 10)
 | 
| 
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
170  | 
  "_MBex"        :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<exists>_:#_./ _)" [0, 0, 10] 10)
 | 
| 
 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 
haftmann 
parents: 
62430 
diff
changeset
 | 
171  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
172  | 
translations  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
173  | 
"\<forall>x\<in>#A. P" \<rightleftharpoons> "CONST Multiset.Ball A (\<lambda>x. P)"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
174  | 
"\<exists>x\<in>#A. P" \<rightleftharpoons> "CONST Multiset.Bex A (\<lambda>x. P)"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
175  | 
|
| 
71917
 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 
nipkow 
parents: 
71398 
diff
changeset
 | 
176  | 
print_translation \<open>  | 
| 
 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 
nipkow 
parents: 
71398 
diff
changeset
 | 
177  | 
[Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\<open>Multiset.Ball\<close> \<^syntax_const>\<open>_MBall\<close>,  | 
| 
 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 
nipkow 
parents: 
71398 
diff
changeset
 | 
178  | 
Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\<open>Multiset.Bex\<close> \<^syntax_const>\<open>_MBex\<close>]  | 
| 
 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 
nipkow 
parents: 
71398 
diff
changeset
 | 
179  | 
\<close> \<comment> \<open>to avoid eta-contraction of body\<close>  | 
| 
 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 
nipkow 
parents: 
71398 
diff
changeset
 | 
180  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
181  | 
lemma count_eq_zero_iff:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
182  | 
"count M x = 0 \<longleftrightarrow> x \<notin># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
183  | 
by (auto simp add: set_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
184  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
185  | 
lemma not_in_iff:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
186  | 
"x \<notin># M \<longleftrightarrow> count M x = 0"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
187  | 
by (auto simp add: count_eq_zero_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
188  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
189  | 
lemma count_greater_zero_iff [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
190  | 
"count M x > 0 \<longleftrightarrow> x \<in># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
191  | 
by (auto simp add: set_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
192  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
193  | 
lemma count_inI:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
194  | 
assumes "count M x = 0 \<Longrightarrow> False"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
195  | 
shows "x \<in># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
196  | 
proof (rule ccontr)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
197  | 
assume "x \<notin># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
198  | 
with assms show False by (simp add: not_in_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
199  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
200  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
201  | 
lemma in_countE:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
202  | 
assumes "x \<in># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
203  | 
obtains n where "count M x = Suc n"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
204  | 
proof -  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
205  | 
from assms have "count M x > 0" by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
206  | 
then obtain n where "count M x = Suc n"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
207  | 
using gr0_conv_Suc by blast  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
208  | 
with that show thesis .  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
209  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
210  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
211  | 
lemma count_greater_eq_Suc_zero_iff [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
212  | 
"count M x \<ge> Suc 0 \<longleftrightarrow> x \<in># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
213  | 
by (simp add: Suc_le_eq)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
214  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
215  | 
lemma count_greater_eq_one_iff [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
216  | 
"count M x \<ge> 1 \<longleftrightarrow> x \<in># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
217  | 
by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
218  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
219  | 
lemma set_mset_empty [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
220  | 
  "set_mset {#} = {}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
221  | 
by (simp add: set_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
222  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
223  | 
lemma set_mset_single:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
224  | 
  "set_mset {#b#} = {b}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
225  | 
by (simp add: set_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
226  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
227  | 
lemma set_mset_eq_empty_iff [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
228  | 
  "set_mset M = {} \<longleftrightarrow> M = {#}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
229  | 
by (auto simp add: multiset_eq_iff count_eq_zero_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
230  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
231  | 
lemma finite_set_mset [iff]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
232  | 
"finite (set_mset M)"  | 
| 73270 | 233  | 
using count [of M] by simp  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
234  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
235  | 
lemma set_mset_add_mset_insert [simp]: \<open>set_mset (add_mset a A) = insert a (set_mset A)\<close>  | 
| 68406 | 236  | 
by (auto simp flip: count_greater_eq_Suc_zero_iff split: if_splits)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
237  | 
|
| 63924 | 238  | 
lemma multiset_nonemptyE [elim]:  | 
239  | 
  assumes "A \<noteq> {#}"
 | 
|
240  | 
obtains x where "x \<in># A"  | 
|
241  | 
proof -  | 
|
242  | 
have "\<exists>x. x \<in># A" by (rule ccontr) (insert assms, auto)  | 
|
243  | 
with that show ?thesis by blast  | 
|
244  | 
qed  | 
|
245  | 
||
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
246  | 
|
| 60500 | 247  | 
subsubsection \<open>Union\<close>  | 
| 10249 | 248  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
249  | 
lemma count_union [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
250  | 
"count (M + N) a = count M a + count N a"  | 
| 
47429
 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 
bulwahn 
parents: 
47308 
diff
changeset
 | 
251  | 
by (simp add: plus_multiset.rep_eq)  | 
| 10249 | 252  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
253  | 
lemma set_mset_union [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
254  | 
"set_mset (M + N) = set_mset M \<union> set_mset N"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
255  | 
by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_union) simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
256  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
257  | 
lemma union_mset_add_mset_left [simp]:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
258  | 
"add_mset a A + B = add_mset a (A + B)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
259  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
260  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
261  | 
lemma union_mset_add_mset_right [simp]:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
262  | 
"A + add_mset a B = add_mset a (A + B)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
263  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
264  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
265  | 
lemma add_mset_add_single: \<open>add_mset a A = A + {#a#}\<close>
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
266  | 
by (subst union_mset_add_mset_right, subst add.comm_neutral) standard  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
267  | 
|
| 10249 | 268  | 
|
| 60500 | 269  | 
subsubsection \<open>Difference\<close>  | 
| 10249 | 270  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
271  | 
instance multiset :: (type) comm_monoid_diff  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
272  | 
by standard (transfer; simp add: fun_eq_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
273  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
274  | 
lemma count_diff [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
275  | 
"count (M - N) a = count M a - count N a"  | 
| 
47429
 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 
bulwahn 
parents: 
47308 
diff
changeset
 | 
276  | 
by (simp add: minus_multiset.rep_eq)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
277  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
278  | 
lemma add_mset_diff_bothsides:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
279  | 
\<open>add_mset a M - add_mset a A = M - A\<close>  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
280  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
281  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
282  | 
lemma in_diff_count:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
283  | 
"a \<in># M - N \<longleftrightarrow> count N a < count M a"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
284  | 
by (simp add: set_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
285  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
286  | 
lemma count_in_diffI:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
287  | 
assumes "\<And>n. count N x = n + count M x \<Longrightarrow> False"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
288  | 
shows "x \<in># M - N"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
289  | 
proof (rule ccontr)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
290  | 
assume "x \<notin># M - N"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
291  | 
then have "count N x = (count N x - count M x) + count M x"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
292  | 
by (simp add: in_diff_count not_less)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
293  | 
with assms show False by auto  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
294  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
295  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
296  | 
lemma in_diff_countE:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
297  | 
assumes "x \<in># M - N"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
298  | 
obtains n where "count M x = Suc n + count N x"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
299  | 
proof -  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
300  | 
from assms have "count M x - count N x > 0" by (simp add: in_diff_count)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
301  | 
then have "count M x > count N x" by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
302  | 
then obtain n where "count M x = Suc n + count N x"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
303  | 
using less_iff_Suc_add by auto  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
304  | 
with that show thesis .  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
305  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
306  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
307  | 
lemma in_diffD:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
308  | 
assumes "a \<in># M - N"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
309  | 
shows "a \<in># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
310  | 
proof -  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
311  | 
have "0 \<le> count N a" by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
312  | 
also from assms have "count N a < count M a"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
313  | 
by (simp add: in_diff_count)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
314  | 
finally show ?thesis by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
315  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
316  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
317  | 
lemma set_mset_diff:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
318  | 
  "set_mset (M - N) = {a. count N a < count M a}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
319  | 
by (simp add: set_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
320  | 
|
| 17161 | 321  | 
lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
 | 
| 52289 | 322  | 
by rule (fact Groups.diff_zero, fact Groups.zero_diff)  | 
| 36903 | 323  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
324  | 
lemma diff_cancel: "A - A = {#}"
 | 
| 52289 | 325  | 
by (fact Groups.diff_cancel)  | 
| 10249 | 326  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
327  | 
lemma diff_union_cancelR: "M + N - N = (M::'a multiset)"  | 
| 52289 | 328  | 
by (fact add_diff_cancel_right')  | 
| 10249 | 329  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
330  | 
lemma diff_union_cancelL: "N + M - N = (M::'a multiset)"  | 
| 52289 | 331  | 
by (fact add_diff_cancel_left')  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
332  | 
|
| 52289 | 333  | 
lemma diff_right_commute:  | 
| 60606 | 334  | 
fixes M N Q :: "'a multiset"  | 
335  | 
shows "M - N - Q = M - Q - N"  | 
|
| 52289 | 336  | 
by (fact diff_right_commute)  | 
337  | 
||
338  | 
lemma diff_add:  | 
|
| 60606 | 339  | 
fixes M N Q :: "'a multiset"  | 
340  | 
shows "M - (N + Q) = M - N - Q"  | 
|
| 52289 | 341  | 
by (rule sym) (fact diff_diff_add)  | 
| 58425 | 342  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
343  | 
lemma insert_DiffM [simp]: "x \<in># M \<Longrightarrow> add_mset x (M - {#x#}) = M"
 | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
344  | 
by (clarsimp simp: multiset_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
345  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
346  | 
lemma insert_DiffM2: "x \<in># M \<Longrightarrow> (M - {#x#}) + {#x#} = M"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
347  | 
by simp  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
348  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
349  | 
lemma diff_union_swap: "a \<noteq> b \<Longrightarrow> add_mset b (M - {#a#}) = add_mset b M - {#a#}"
 | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
350  | 
by (auto simp add: multiset_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
351  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
352  | 
lemma diff_add_mset_swap [simp]: "b \<notin># A \<Longrightarrow> add_mset b M - A = add_mset b (M - A)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
353  | 
by (auto simp add: multiset_eq_iff simp: not_in_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
354  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
355  | 
lemma diff_union_swap2 [simp]: "y \<in># M \<Longrightarrow> add_mset x M - {#y#} = add_mset x (M - {#y#})"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
356  | 
by (metis add_mset_diff_bothsides diff_union_swap diff_zero insert_DiffM)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
357  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
358  | 
lemma diff_diff_add_mset [simp]: "(M::'a multiset) - N - P = M - (N + P)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
359  | 
by (rule diff_diff_add)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
360  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
361  | 
lemma diff_union_single_conv:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
362  | 
  "a \<in># J \<Longrightarrow> I + J - {#a#} = I + (J - {#a#})"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
363  | 
by (simp add: multiset_eq_iff Suc_le_eq)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
364  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
365  | 
lemma mset_add [elim?]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
366  | 
assumes "a \<in># A"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
367  | 
obtains B where "A = add_mset a B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
368  | 
proof -  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
369  | 
  from assms have "A = add_mset a (A - {#a#})"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
370  | 
by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
371  | 
with that show thesis .  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
372  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
373  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
374  | 
lemma union_iff:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
375  | 
"a \<in># A + B \<longleftrightarrow> a \<in># A \<or> a \<in># B"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
376  | 
by auto  | 
| 
26143
 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 
bulwahn 
parents: 
26033 
diff
changeset
 | 
377  | 
|
| 
77987
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
378  | 
lemma count_minus_inter_lt_count_minus_inter_iff:  | 
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
379  | 
"count (M2 - M1) y < count (M1 - M2) y \<longleftrightarrow> y \<in># M1 - M2"  | 
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
380  | 
by (meson count_greater_zero_iff gr_implies_not_zero in_diff_count leI order.strict_trans2  | 
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
381  | 
order_less_asym)  | 
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
382  | 
|
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
383  | 
lemma minus_inter_eq_minus_inter_iff:  | 
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
384  | 
"(M1 - M2) = (M2 - M1) \<longleftrightarrow> set_mset (M1 - M2) = set_mset (M2 - M1)"  | 
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
385  | 
by (metis add.commute count_diff count_eq_zero_iff diff_add_zero in_diff_countE multiset_eq_iff)  | 
| 
 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 
desharna 
parents: 
77832 
diff
changeset
 | 
386  | 
|
| 10249 | 387  | 
|
| 66425 | 388  | 
subsubsection \<open>Min and Max\<close>  | 
389  | 
||
390  | 
abbreviation Min_mset :: "'a::linorder multiset \<Rightarrow> 'a" where  | 
|
391  | 
"Min_mset m \<equiv> Min (set_mset m)"  | 
|
392  | 
||
393  | 
abbreviation Max_mset :: "'a::linorder multiset \<Rightarrow> 'a" where  | 
|
394  | 
"Max_mset m \<equiv> Max (set_mset m)"  | 
|
395  | 
||
396  | 
||
| 60500 | 397  | 
subsubsection \<open>Equality of multisets\<close>  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
398  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
399  | 
lemma single_eq_single [simp]: "{#a#} = {#b#} \<longleftrightarrow> a = b"
 | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
400  | 
by (auto simp add: multiset_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
401  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
402  | 
lemma union_eq_empty [iff]: "M + N = {#} \<longleftrightarrow> M = {#} \<and> N = {#}"
 | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
403  | 
by (auto simp add: multiset_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
404  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
405  | 
lemma empty_eq_union [iff]: "{#} = M + N \<longleftrightarrow> M = {#} \<and> N = {#}"
 | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
406  | 
by (auto simp add: multiset_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
407  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
408  | 
lemma multi_self_add_other_not_self [simp]: "M = add_mset x M \<longleftrightarrow> False"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
409  | 
by (auto simp add: multiset_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
410  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
411  | 
lemma add_mset_remove_trivial [simp]: \<open>add_mset x M - {#x#} = M\<close>
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
412  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
413  | 
|
| 60606 | 414  | 
lemma diff_single_trivial: "\<not> x \<in># M \<Longrightarrow> M - {#x#} = M"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
415  | 
by (auto simp add: multiset_eq_iff not_in_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
416  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
417  | 
lemma diff_single_eq_union: "x \<in># M \<Longrightarrow> M - {#x#} = N \<longleftrightarrow> M = add_mset x N"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
418  | 
by auto  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
419  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
420  | 
lemma union_single_eq_diff: "add_mset x M = N \<Longrightarrow> M = N - {#x#}"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
421  | 
unfolding add_mset_add_single[of _ M] by (fact add_implies_diff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
422  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
423  | 
lemma union_single_eq_member: "add_mset x M = N \<Longrightarrow> x \<in># N"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
424  | 
by auto  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
425  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
426  | 
lemma add_mset_remove_trivial_If:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
427  | 
  "add_mset a (N - {#a#}) = (if a \<in># N then N else add_mset a N)"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
428  | 
by (simp add: diff_single_trivial)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
429  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
430  | 
lemma add_mset_remove_trivial_eq: \<open>N = add_mset a (N - {#a#}) \<longleftrightarrow> a \<in># N\<close>
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
431  | 
by (auto simp: add_mset_remove_trivial_If)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
432  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
433  | 
lemma union_is_single:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
434  | 
  "M + N = {#a#} \<longleftrightarrow> M = {#a#} \<and> N = {#} \<or> M = {#} \<and> N = {#a#}"
 | 
| 60606 | 435  | 
(is "?lhs = ?rhs")  | 
| 46730 | 436  | 
proof  | 
| 60606 | 437  | 
show ?lhs if ?rhs using that by auto  | 
438  | 
show ?rhs if ?lhs  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
439  | 
by (metis Multiset.diff_cancel add.commute add_diff_cancel_left' diff_add_zero diff_single_trivial insert_DiffM that)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
440  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
441  | 
|
| 60606 | 442  | 
lemma single_is_union: "{#a#} = M + N \<longleftrightarrow> {#a#} = M \<and> N = {#} \<or> M = {#} \<and> {#a#} = N"
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
443  | 
  by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single)
 | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
444  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
445  | 
lemma add_eq_conv_diff:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
446  | 
  "add_mset a M = add_mset b N \<longleftrightarrow> M = N \<and> a = b \<or> M = add_mset b (N - {#a#}) \<and> N = add_mset a (M - {#b#})"
 | 
| 60606 | 447  | 
(is "?lhs \<longleftrightarrow> ?rhs")  | 
| 
44890
 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 
nipkow 
parents: 
44339 
diff
changeset
 | 
448  | 
(* shorter: by (simp add: multiset_eq_iff) fastforce *)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
449  | 
proof  | 
| 60606 | 450  | 
show ?lhs if ?rhs  | 
451  | 
using that  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
452  | 
by (auto simp add: add_mset_commute[of a b])  | 
| 60606 | 453  | 
show ?rhs if ?lhs  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
454  | 
proof (cases "a = b")  | 
| 60500 | 455  | 
case True with \<open>?lhs\<close> show ?thesis by simp  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
456  | 
next  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
457  | 
case False  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
458  | 
from \<open>?lhs\<close> have "a \<in># add_mset b N" by (rule union_single_eq_member)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
459  | 
with False have "a \<in># N" by auto  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
460  | 
    moreover from \<open>?lhs\<close> have "M = add_mset b N - {#a#}" by (rule union_single_eq_diff)
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
461  | 
moreover note False  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
462  | 
    ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"])
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
463  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
464  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
465  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
466  | 
lemma add_mset_eq_single [iff]: "add_mset b M = {#a#} \<longleftrightarrow> b = a \<and> M = {#}"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
467  | 
by (auto simp: add_eq_conv_diff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
468  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
469  | 
lemma single_eq_add_mset [iff]: "{#a#} = add_mset b M \<longleftrightarrow> b = a \<and> M = {#}"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
470  | 
by (auto simp: add_eq_conv_diff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
471  | 
|
| 58425 | 472  | 
lemma insert_noteq_member:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
473  | 
assumes BC: "add_mset b B = add_mset c C"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
474  | 
and bnotc: "b \<noteq> c"  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
475  | 
shows "c \<in># B"  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
476  | 
proof -  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
477  | 
have "c \<in># add_mset c C" by simp  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
478  | 
  have nc: "\<not> c \<in># {#b#}" using bnotc by simp
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
479  | 
then have "c \<in># add_mset b B" using BC by simp  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
480  | 
then show "c \<in># B" using nc by simp  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
481  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
482  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
483  | 
lemma add_eq_conv_ex:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
484  | 
"(add_mset a M = add_mset b N) =  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
485  | 
(M = N \<and> a = b \<or> (\<exists>K. M = add_mset b K \<and> N = add_mset a K))"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
486  | 
by (auto simp add: add_eq_conv_diff)  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
487  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
488  | 
lemma multi_member_split: "x \<in># M \<Longrightarrow> \<exists>A. M = add_mset x A"  | 
| 60678 | 489  | 
  by (rule exI [where x = "M - {#x#}"]) simp
 | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
490  | 
|
| 58425 | 491  | 
lemma multiset_add_sub_el_shuffle:  | 
| 60606 | 492  | 
assumes "c \<in># B"  | 
493  | 
and "b \<noteq> c"  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
494  | 
  shows "add_mset b (B - {#c#}) = add_mset b B - {#c#}"
 | 
| 58098 | 495  | 
proof -  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
496  | 
from \<open>c \<in># B\<close> obtain A where B: "B = add_mset c A"  | 
| 58098 | 497  | 
by (blast dest: multi_member_split)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
498  | 
  have "add_mset b A = add_mset c (add_mset b A) - {#c#}" by simp
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
499  | 
  then have "add_mset b A = add_mset b (add_mset c A) - {#c#}"
 | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
500  | 
by (simp add: \<open>b \<noteq> c\<close>)  | 
| 58098 | 501  | 
then show ?thesis using B by simp  | 
502  | 
qed  | 
|
503  | 
||
| 64418 | 504  | 
lemma add_mset_eq_singleton_iff[iff]:  | 
505  | 
  "add_mset x M = {#y#} \<longleftrightarrow> M = {#} \<and> x = y"
 | 
|
506  | 
by auto  | 
|
507  | 
||
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
508  | 
|
| 60500 | 509  | 
subsubsection \<open>Pointwise ordering induced by count\<close>  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
510  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
511  | 
definition subseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subseteq>#" 50)  | 
| 65466 | 512  | 
where "A \<subseteq># B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)"  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
513  | 
|
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
514  | 
definition subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subset>#" 50)  | 
| 65466 | 515  | 
where "A \<subset># B \<longleftrightarrow> A \<subseteq># B \<and> A \<noteq> B"  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
516  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
517  | 
abbreviation (input) supseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<supseteq>#" 50)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
518  | 
where "supseteq_mset A B \<equiv> B \<subseteq># A"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
519  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
520  | 
abbreviation (input) supset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<supset>#" 50)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
521  | 
where "supset_mset A B \<equiv> B \<subset># A"  | 
| 
62208
 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 
blanchet 
parents: 
62082 
diff
changeset
 | 
522  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
523  | 
notation (input)  | 
| 
62208
 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 
blanchet 
parents: 
62082 
diff
changeset
 | 
524  | 
subseteq_mset (infix "\<le>#" 50) and  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
525  | 
supseteq_mset (infix "\<ge>#" 50)  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
526  | 
|
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
527  | 
notation (ASCII)  | 
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
528  | 
subseteq_mset (infix "<=#" 50) and  | 
| 
62208
 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 
blanchet 
parents: 
62082 
diff
changeset
 | 
529  | 
subset_mset (infix "<#" 50) and  | 
| 
 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 
blanchet 
parents: 
62082 
diff
changeset
 | 
530  | 
supseteq_mset (infix ">=#" 50) and  | 
| 
 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 
blanchet 
parents: 
62082 
diff
changeset
 | 
531  | 
supset_mset (infix ">#" 50)  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
532  | 
|
| 73411 | 533  | 
global_interpretation subset_mset: ordering \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close>  | 
534  | 
by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order.trans order.antisym)  | 
|
535  | 
||
| 73451 | 536  | 
interpretation subset_mset: ordered_ab_semigroup_add_imp_le \<open>(+)\<close> \<open>(-)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close>  | 
| 60678 | 537  | 
by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym)  | 
| 
64585
 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 
haftmann 
parents: 
64531 
diff
changeset
 | 
538  | 
\<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
539  | 
|
| 67398 | 540  | 
interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "(+)" 0 "(-)" "(\<subseteq>#)" "(\<subset>#)"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
541  | 
by standard  | 
| 
64585
 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 
haftmann 
parents: 
64531 
diff
changeset
 | 
542  | 
\<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
543  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
544  | 
lemma mset_subset_eqI:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
545  | 
"(\<And>a. count A a \<le> count B a) \<Longrightarrow> A \<subseteq># B"  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
546  | 
by (simp add: subseteq_mset_def)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
547  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
548  | 
lemma mset_subset_eq_count:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
549  | 
"A \<subseteq># B \<Longrightarrow> count A a \<le> count B a"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
550  | 
by (simp add: subseteq_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
551  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
552  | 
lemma mset_subset_eq_exists_conv: "(A::'a multiset) \<subseteq># B \<longleftrightarrow> (\<exists>C. B = A + C)"  | 
| 60678 | 553  | 
unfolding subseteq_mset_def  | 
554  | 
apply (rule iffI)  | 
|
555  | 
apply (rule exI [where x = "B - A"])  | 
|
556  | 
apply (auto intro: multiset_eq_iff [THEN iffD2])  | 
|
557  | 
done  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
558  | 
|
| 67398 | 559  | 
interpretation subset_mset: ordered_cancel_comm_monoid_diff "(+)" 0 "(\<subseteq>#)" "(\<subset>#)" "(-)"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
560  | 
by standard (simp, fact mset_subset_eq_exists_conv)  | 
| 
64585
 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 
haftmann 
parents: 
64531 
diff
changeset
 | 
561  | 
\<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
562  | 
|
| 
64017
 
6e7bf7678518
more multiset simp rules
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63924 
diff
changeset
 | 
563  | 
declare subset_mset.add_diff_assoc[simp] subset_mset.add_diff_assoc2[simp]  | 
| 
 
6e7bf7678518
more multiset simp rules
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63924 
diff
changeset
 | 
564  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
565  | 
lemma mset_subset_eq_mono_add_right_cancel: "(A::'a multiset) + C \<subseteq># B + C \<longleftrightarrow> A \<subseteq># B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
566  | 
by (fact subset_mset.add_le_cancel_right)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
567  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
568  | 
lemma mset_subset_eq_mono_add_left_cancel: "C + (A::'a multiset) \<subseteq># C + B \<longleftrightarrow> A \<subseteq># B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
569  | 
by (fact subset_mset.add_le_cancel_left)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
570  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
571  | 
lemma mset_subset_eq_mono_add: "(A::'a multiset) \<subseteq># B \<Longrightarrow> C \<subseteq># D \<Longrightarrow> A + C \<subseteq># B + D"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
572  | 
by (fact subset_mset.add_mono)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
573  | 
|
| 
63560
 
3e3097ac37d1
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63547 
diff
changeset
 | 
574  | 
lemma mset_subset_eq_add_left: "(A::'a multiset) \<subseteq># A + B"  | 
| 
 
3e3097ac37d1
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63547 
diff
changeset
 | 
575  | 
by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
576  | 
|
| 
63560
 
3e3097ac37d1
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63547 
diff
changeset
 | 
577  | 
lemma mset_subset_eq_add_right: "B \<subseteq># (A::'a multiset) + B"  | 
| 
 
3e3097ac37d1
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63547 
diff
changeset
 | 
578  | 
by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
579  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
580  | 
lemma single_subset_iff [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
581  | 
  "{#a#} \<subseteq># M \<longleftrightarrow> a \<in># M"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
582  | 
by (auto simp add: subseteq_mset_def Suc_le_eq)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
583  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
584  | 
lemma mset_subset_eq_single: "a \<in># B \<Longrightarrow> {#a#} \<subseteq># B"
 | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
585  | 
by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
586  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
587  | 
lemma mset_subset_eq_add_mset_cancel: \<open>add_mset a A \<subseteq># add_mset a B \<longleftrightarrow> A \<subseteq># B\<close>  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
588  | 
unfolding add_mset_add_single[of _ A] add_mset_add_single[of _ B]  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
589  | 
by (rule mset_subset_eq_mono_add_right_cancel)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
590  | 
|
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
591  | 
lemma multiset_diff_union_assoc:  | 
| 60606 | 592  | 
fixes A B C D :: "'a multiset"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
593  | 
shows "C \<subseteq># B \<Longrightarrow> A + B - C = A + (B - C)"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
594  | 
by (fact subset_mset.diff_add_assoc)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
595  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
596  | 
lemma mset_subset_eq_multiset_union_diff_commute:  | 
| 60606 | 597  | 
fixes A B C D :: "'a multiset"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
598  | 
shows "B \<subseteq># A \<Longrightarrow> A - B + C = A + C - B"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
599  | 
by (fact subset_mset.add_diff_assoc2)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
600  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
601  | 
lemma diff_subset_eq_self[simp]:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
602  | 
"(M::'a multiset) - N \<subseteq># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
603  | 
by (simp add: subseteq_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
604  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
605  | 
lemma mset_subset_eqD:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
606  | 
assumes "A \<subseteq># B" and "x \<in># A"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
607  | 
shows "x \<in># B"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
608  | 
proof -  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
609  | 
from \<open>x \<in># A\<close> have "count A x > 0" by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
610  | 
also from \<open>A \<subseteq># B\<close> have "count A x \<le> count B x"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
611  | 
by (simp add: subseteq_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
612  | 
finally show ?thesis by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
613  | 
qed  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
614  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
615  | 
lemma mset_subsetD:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
616  | 
"A \<subset># B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
617  | 
by (auto intro: mset_subset_eqD [of A])  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
618  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
619  | 
lemma set_mset_mono:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
620  | 
"A \<subseteq># B \<Longrightarrow> set_mset A \<subseteq> set_mset B"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
621  | 
by (metis mset_subset_eqD subsetI)  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
622  | 
|
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
623  | 
lemma mset_subset_eq_insertD:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
624  | 
"add_mset x A \<subseteq># B \<Longrightarrow> x \<in># B \<and> A \<subset># B"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
625  | 
apply (rule conjI)  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
626  | 
apply (simp add: mset_subset_eqD)  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
627  | 
apply (clarsimp simp: subset_mset_def subseteq_mset_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
628  | 
apply safe  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
629  | 
apply (erule_tac x = a in allE)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
630  | 
apply (auto split: if_split_asm)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
631  | 
done  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
632  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
633  | 
lemma mset_subset_insertD:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
634  | 
"add_mset x A \<subset># B \<Longrightarrow> x \<in># B \<and> A \<subset># B"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
635  | 
by (rule mset_subset_eq_insertD) simp  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
636  | 
|
| 63831 | 637  | 
lemma mset_subset_of_empty[simp]: "A \<subset># {#} \<longleftrightarrow> False"
 | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
638  | 
by (simp only: subset_mset.not_less_zero)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
639  | 
|
| 64587 | 640  | 
lemma empty_subset_add_mset[simp]: "{#} \<subset># add_mset x M"
 | 
641  | 
by (auto intro: subset_mset.gr_zeroI)  | 
|
| 63831 | 642  | 
|
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
643  | 
lemma empty_le: "{#} \<subseteq># A"
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
644  | 
by (fact subset_mset.zero_le)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
645  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
646  | 
lemma insert_subset_eq_iff:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
647  | 
  "add_mset a A \<subseteq># B \<longleftrightarrow> a \<in># B \<and> A \<subseteq># B - {#a#}"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
648  | 
using le_diff_conv2 [of "Suc 0" "count B a" "count A a"]  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
649  | 
apply (auto simp add: subseteq_mset_def not_in_iff Suc_le_eq)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
650  | 
apply (rule ccontr)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
651  | 
apply (auto simp add: not_in_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
652  | 
done  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
653  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
654  | 
lemma insert_union_subset_iff:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
655  | 
  "add_mset a A \<subset># B \<longleftrightarrow> a \<in># B \<and> A \<subset># B - {#a#}"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
656  | 
by (auto simp add: insert_subset_eq_iff subset_mset_def)  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
657  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
658  | 
lemma subset_eq_diff_conv:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
659  | 
"A - C \<subseteq># B \<longleftrightarrow> A \<subseteq># B + C"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
660  | 
by (simp add: subseteq_mset_def le_diff_conv)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
661  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
662  | 
lemma multi_psub_of_add_self [simp]: "A \<subset># add_mset x A"  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
663  | 
by (auto simp: subset_mset_def subseteq_mset_def)  | 
| 
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
664  | 
|
| 64076 | 665  | 
lemma multi_psub_self: "A \<subset># A = False"  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
666  | 
by simp  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
667  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
668  | 
lemma mset_subset_add_mset [simp]: "add_mset x N \<subset># add_mset x M \<longleftrightarrow> N \<subset># M"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
669  | 
unfolding add_mset_add_single[of _ N] add_mset_add_single[of _ M]  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
670  | 
by (fact subset_mset.add_less_cancel_right)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
671  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
672  | 
lemma mset_subset_diff_self: "c \<in># B \<Longrightarrow> B - {#c#} \<subset># B"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
673  | 
by (auto simp: subset_mset_def elim: mset_add)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
674  | 
|
| 64077 | 675  | 
lemma Diff_eq_empty_iff_mset: "A - B = {#} \<longleftrightarrow> A \<subseteq># B"
 | 
676  | 
by (auto simp: multiset_eq_iff subseteq_mset_def)  | 
|
677  | 
||
| 64418 | 678  | 
lemma add_mset_subseteq_single_iff[iff]: "add_mset a M \<subseteq># {#b#} \<longleftrightarrow> M = {#} \<and> a = b"
 | 
679  | 
proof  | 
|
680  | 
  assume A: "add_mset a M \<subseteq># {#b#}"
 | 
|
681  | 
then have \<open>a = b\<close>  | 
|
682  | 
by (auto dest: mset_subset_eq_insertD)  | 
|
683  | 
  then show "M={#} \<and> a=b"
 | 
|
684  | 
using A by (simp add: mset_subset_eq_add_mset_cancel)  | 
|
685  | 
qed simp  | 
|
686  | 
||
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
687  | 
|
| 64076 | 688  | 
subsubsection \<open>Intersection and bounded union\<close>  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
689  | 
|
| 73393 | 690  | 
definition inter_mset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> (infixl \<open>\<inter>#\<close> 70)  | 
691  | 
where \<open>A \<inter># B = A - (A - B)\<close>  | 
|
692  | 
||
| 73451 | 693  | 
lemma count_inter_mset [simp]:  | 
694  | 
\<open>count (A \<inter># B) x = min (count A x) (count B x)\<close>  | 
|
695  | 
by (simp add: inter_mset_def)  | 
|
696  | 
||
697  | 
(*global_interpretation subset_mset: semilattice_order \<open>(\<inter>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close>  | 
|
698  | 
by standard (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def min_def)*)  | 
|
699  | 
||
| 73393 | 700  | 
interpretation subset_mset: semilattice_inf \<open>(\<inter>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close>  | 
| 73451 | 701  | 
by standard (simp_all add: multiset_eq_iff subseteq_mset_def)  | 
702  | 
\<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
703  | 
|
| 73393 | 704  | 
definition union_mset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> (infixl \<open>\<union>#\<close> 70)  | 
705  | 
where \<open>A \<union># B = A + (B - A)\<close>  | 
|
706  | 
||
| 73451 | 707  | 
lemma count_union_mset [simp]:  | 
708  | 
\<open>count (A \<union># B) x = max (count A x) (count B x)\<close>  | 
|
709  | 
by (simp add: union_mset_def)  | 
|
710  | 
||
711  | 
global_interpretation subset_mset: semilattice_neutr_order \<open>(\<union>#)\<close> \<open>{#}\<close> \<open>(\<supseteq>#)\<close> \<open>(\<supset>#)\<close>
 | 
|
712  | 
apply standard  | 
|
713  | 
apply (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def max_def)  | 
|
714  | 
apply (auto simp add: le_antisym dest: sym)  | 
|
715  | 
apply (metis nat_le_linear)+  | 
|
716  | 
done  | 
|
717  | 
||
| 73393 | 718  | 
interpretation subset_mset: semilattice_sup \<open>(\<union>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close>  | 
| 64076 | 719  | 
proof -  | 
720  | 
have [simp]: "m \<le> n \<Longrightarrow> q \<le> n \<Longrightarrow> m + (q - m) \<le> n" for m n q :: nat  | 
|
721  | 
by arith  | 
|
| 67398 | 722  | 
show "class.semilattice_sup (\<union>#) (\<subseteq>#) (\<subset>#)"  | 
| 73393 | 723  | 
by standard (auto simp add: union_mset_def subseteq_mset_def)  | 
| 
64585
 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 
haftmann 
parents: 
64531 
diff
changeset
 | 
724  | 
qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
| 64076 | 725  | 
|
| 67398 | 726  | 
interpretation subset_mset: bounded_lattice_bot "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)"  | 
727  | 
  "(\<union>#)" "{#}"
 | 
|
| 64076 | 728  | 
by standard auto  | 
| 
64585
 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 
haftmann 
parents: 
64531 
diff
changeset
 | 
729  | 
\<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
| 64076 | 730  | 
|
731  | 
||
732  | 
subsubsection \<open>Additional intersection facts\<close>  | 
|
733  | 
||
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
734  | 
lemma set_mset_inter [simp]:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
735  | 
"set_mset (A \<inter># B) = set_mset A \<inter> set_mset B"  | 
| 73393 | 736  | 
by (simp only: set_mset_def) auto  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
737  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
738  | 
lemma diff_intersect_left_idem [simp]:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
739  | 
"M - M \<inter># N = M - N"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
740  | 
by (simp add: multiset_eq_iff min_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
741  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
742  | 
lemma diff_intersect_right_idem [simp]:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
743  | 
"M - N \<inter># M = M - N"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
744  | 
by (simp add: multiset_eq_iff min_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
745  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
746  | 
lemma multiset_inter_single[simp]: "a \<noteq> b \<Longrightarrow> {#a#} \<inter># {#b#} = {#}"
 | 
| 46730 | 747  | 
by (rule multiset_eqI) auto  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
748  | 
|
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
749  | 
lemma multiset_union_diff_commute:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
750  | 
  assumes "B \<inter># C = {#}"
 | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
751  | 
shows "A + B - C = A - C + B"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
752  | 
proof (rule multiset_eqI)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
753  | 
fix x  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
754  | 
from assms have "min (count B x) (count C x) = 0"  | 
| 46730 | 755  | 
by (auto simp add: multiset_eq_iff)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
756  | 
then have "count B x = 0 \<or> count C x = 0"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
757  | 
unfolding min_def by (auto split: if_splits)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
758  | 
then show "count (A + B - C) x = count (A - C + B) x"  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
759  | 
by auto  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
760  | 
qed  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
761  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
762  | 
lemma disjunct_not_in:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
763  | 
  "A \<inter># B = {#} \<longleftrightarrow> (\<forall>a. a \<notin># A \<or> a \<notin># B)" (is "?P \<longleftrightarrow> ?Q")
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
764  | 
proof  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
765  | 
assume ?P  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
766  | 
show ?Q  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
767  | 
proof  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
768  | 
fix a  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
769  | 
from \<open>?P\<close> have "min (count A a) (count B a) = 0"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
770  | 
by (simp add: multiset_eq_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
771  | 
then have "count A a = 0 \<or> count B a = 0"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
772  | 
by (cases "count A a \<le> count B a") (simp_all add: min_def)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
773  | 
then show "a \<notin># A \<or> a \<notin># B"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
774  | 
by (simp add: not_in_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
775  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
776  | 
next  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
777  | 
assume ?Q  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
778  | 
show ?P  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
779  | 
proof (rule multiset_eqI)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
780  | 
fix a  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
781  | 
from \<open>?Q\<close> have "count A a = 0 \<or> count B a = 0"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
782  | 
by (auto simp add: not_in_iff)  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
783  | 
    then show "count (A \<inter># B) a = count {#} a"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
784  | 
by auto  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
785  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
786  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
787  | 
|
| 64077 | 788  | 
lemma inter_mset_empty_distrib_right: "A \<inter># (B + C) = {#} \<longleftrightarrow> A \<inter># B = {#} \<and> A \<inter># C = {#}"
 | 
789  | 
by (meson disjunct_not_in union_iff)  | 
|
790  | 
||
791  | 
lemma inter_mset_empty_distrib_left: "(A + B) \<inter># C = {#} \<longleftrightarrow> A \<inter># C = {#} \<and> B \<inter># C = {#}"
 | 
|
792  | 
by (meson disjunct_not_in union_iff)  | 
|
793  | 
||
| 73393 | 794  | 
lemma add_mset_inter_add_mset [simp]:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
795  | 
"add_mset a A \<inter># add_mset a B = add_mset a (A \<inter># B)"  | 
| 73393 | 796  | 
by (rule multiset_eqI) simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
797  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
798  | 
lemma add_mset_disjoint [simp]:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
799  | 
  "add_mset a A \<inter># B = {#} \<longleftrightarrow> a \<notin># B \<and> A \<inter># B = {#}"
 | 
| 
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
800  | 
  "{#} = add_mset a A \<inter># B \<longleftrightarrow> a \<notin># B \<and> {#} = A \<inter># B"
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
801  | 
by (auto simp: disjunct_not_in)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
802  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
803  | 
lemma disjoint_add_mset [simp]:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
804  | 
  "B \<inter># add_mset a A = {#} \<longleftrightarrow> a \<notin># B \<and> B \<inter># A = {#}"
 | 
| 
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
805  | 
  "{#} = A \<inter># add_mset b B \<longleftrightarrow> b \<notin># A \<and> {#} = A \<inter># B"
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
806  | 
by (auto simp: disjunct_not_in)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
807  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
808  | 
lemma inter_add_left1: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) \<inter># N = M \<inter># N"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
809  | 
by (simp add: multiset_eq_iff not_in_iff)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
810  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
811  | 
lemma inter_add_left2: "x \<in># N \<Longrightarrow> (add_mset x M) \<inter># N = add_mset x (M \<inter># (N - {#x#}))"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
812  | 
by (auto simp add: multiset_eq_iff elim: mset_add)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
813  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
814  | 
lemma inter_add_right1: "\<not> x \<in># N \<Longrightarrow> N \<inter># (add_mset x M) = N \<inter># M"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
815  | 
by (simp add: multiset_eq_iff not_in_iff)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
816  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
817  | 
lemma inter_add_right2: "x \<in># N \<Longrightarrow> N \<inter># (add_mset x M) = add_mset x ((N - {#x#}) \<inter># M)"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
818  | 
by (auto simp add: multiset_eq_iff elim: mset_add)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
819  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
820  | 
lemma disjunct_set_mset_diff:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
821  | 
  assumes "M \<inter># N = {#}"
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
822  | 
shows "set_mset (M - N) = set_mset M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
823  | 
proof (rule set_eqI)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
824  | 
fix a  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
825  | 
from assms have "a \<notin># M \<or> a \<notin># N"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
826  | 
by (simp add: disjunct_not_in)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
827  | 
then show "a \<in># M - N \<longleftrightarrow> a \<in># M"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
828  | 
by (auto dest: in_diffD) (simp add: in_diff_count not_in_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
829  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
830  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
831  | 
lemma at_most_one_mset_mset_diff:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
832  | 
  assumes "a \<notin># M - {#a#}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
833  | 
  shows "set_mset (M - {#a#}) = set_mset M - {a}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
834  | 
using assms by (auto simp add: not_in_iff in_diff_count set_eq_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
835  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
836  | 
lemma more_than_one_mset_mset_diff:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
837  | 
  assumes "a \<in># M - {#a#}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
838  | 
  shows "set_mset (M - {#a#}) = set_mset M"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
839  | 
proof (rule set_eqI)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
840  | 
fix b  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
841  | 
have "Suc 0 < count M b \<Longrightarrow> count M b > 0" by arith  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
842  | 
  then show "b \<in># M - {#a#} \<longleftrightarrow> b \<in># M"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
843  | 
using assms by (auto simp add: in_diff_count)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
844  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
845  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
846  | 
lemma inter_iff:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
847  | 
"a \<in># A \<inter># B \<longleftrightarrow> a \<in># A \<and> a \<in># B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
848  | 
by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
849  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
850  | 
lemma inter_union_distrib_left:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
851  | 
"A \<inter># B + C = (A + C) \<inter># (B + C)"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
852  | 
by (simp add: multiset_eq_iff min_add_distrib_left)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
853  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
854  | 
lemma inter_union_distrib_right:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
855  | 
"C + A \<inter># B = (C + A) \<inter># (C + B)"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
856  | 
using inter_union_distrib_left [of A B C] by (simp add: ac_simps)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
857  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
858  | 
lemma inter_subset_eq_union:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
859  | 
"A \<inter># B \<subseteq># A + B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
860  | 
by (auto simp add: subseteq_mset_def)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
861  | 
|
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
862  | 
|
| 64076 | 863  | 
subsubsection \<open>Additional bounded union facts\<close>  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
864  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
865  | 
lemma set_mset_sup [simp]:  | 
| 73393 | 866  | 
\<open>set_mset (A \<union># B) = set_mset A \<union> set_mset B\<close>  | 
867  | 
by (simp only: set_mset_def) (auto simp add: less_max_iff_disj)  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
868  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
869  | 
lemma sup_union_left1 [simp]: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) \<union># N = add_mset x (M \<union># N)"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
870  | 
by (simp add: multiset_eq_iff not_in_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
871  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
872  | 
lemma sup_union_left2: "x \<in># N \<Longrightarrow> (add_mset x M) \<union># N = add_mset x (M \<union># (N - {#x#}))"
 | 
| 51623 | 873  | 
by (simp add: multiset_eq_iff)  | 
874  | 
||
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
875  | 
lemma sup_union_right1 [simp]: "\<not> x \<in># N \<Longrightarrow> N \<union># (add_mset x M) = add_mset x (N \<union># M)"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
876  | 
by (simp add: multiset_eq_iff not_in_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
877  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
878  | 
lemma sup_union_right2: "x \<in># N \<Longrightarrow> N \<union># (add_mset x M) = add_mset x ((N - {#x#}) \<union># M)"
 | 
| 51623 | 879  | 
by (simp add: multiset_eq_iff)  | 
880  | 
||
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
881  | 
lemma sup_union_distrib_left:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
882  | 
"A \<union># B + C = (A + C) \<union># (B + C)"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
883  | 
by (simp add: multiset_eq_iff max_add_distrib_left)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
884  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
885  | 
lemma union_sup_distrib_right:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
886  | 
"C + A \<union># B = (C + A) \<union># (C + B)"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
887  | 
using sup_union_distrib_left [of A B C] by (simp add: ac_simps)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
888  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
889  | 
lemma union_diff_inter_eq_sup:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
890  | 
"A + B - A \<inter># B = A \<union># B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
891  | 
by (auto simp add: multiset_eq_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
892  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
893  | 
lemma union_diff_sup_eq_inter:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
894  | 
"A + B - A \<union># B = A \<inter># B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
895  | 
by (auto simp add: multiset_eq_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
896  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
897  | 
lemma add_mset_union:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
898  | 
\<open>add_mset a A \<union># add_mset a B = add_mset a (A \<union># B)\<close>  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
899  | 
by (auto simp: multiset_eq_iff max_def)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
900  | 
|
| 51623 | 901  | 
|
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
902  | 
subsection \<open>Replicate and repeat operations\<close>  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
903  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
904  | 
definition replicate_mset :: "nat \<Rightarrow> 'a \<Rightarrow> 'a multiset" where  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
905  | 
  "replicate_mset n x = (add_mset x ^^ n) {#}"
 | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
906  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
907  | 
lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}"
 | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
908  | 
unfolding replicate_mset_def by simp  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
909  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
910  | 
lemma replicate_mset_Suc [simp]: "replicate_mset (Suc n) x = add_mset x (replicate_mset n x)"  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
911  | 
unfolding replicate_mset_def by (induct n) (auto intro: add.commute)  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
912  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
913  | 
lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)"  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
914  | 
unfolding replicate_mset_def by (induct n) auto  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
915  | 
|
| 73393 | 916  | 
lift_definition repeat_mset :: \<open>nat \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close>  | 
917  | 
is \<open>\<lambda>n M a. n * M a\<close> by simp  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
918  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
919  | 
lemma count_repeat_mset [simp]: "count (repeat_mset i A) a = i * count A a"  | 
| 73393 | 920  | 
by transfer rule  | 
921  | 
||
922  | 
lemma repeat_mset_0 [simp]:  | 
|
923  | 
  \<open>repeat_mset 0 M = {#}\<close>
 | 
|
924  | 
by transfer simp  | 
|
925  | 
||
926  | 
lemma repeat_mset_Suc [simp]:  | 
|
927  | 
\<open>repeat_mset (Suc n) M = M + repeat_mset n M\<close>  | 
|
928  | 
by transfer simp  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
929  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
930  | 
lemma repeat_mset_right [simp]: "repeat_mset a (repeat_mset b A) = repeat_mset (a * b) A"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
931  | 
by (auto simp: multiset_eq_iff left_diff_distrib')  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
932  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
933  | 
lemma left_diff_repeat_mset_distrib': \<open>repeat_mset (i - j) u = repeat_mset i u - repeat_mset j u\<close>  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
934  | 
by (auto simp: multiset_eq_iff left_diff_distrib')  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
935  | 
|
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
936  | 
lemma left_add_mult_distrib_mset:  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
937  | 
"repeat_mset i u + (repeat_mset j u + k) = repeat_mset (i+j) u + k"  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
938  | 
by (auto simp: multiset_eq_iff add_mult_distrib)  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
939  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
940  | 
lemma repeat_mset_distrib:  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
941  | 
"repeat_mset (m + n) A = repeat_mset m A + repeat_mset n A"  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
942  | 
by (auto simp: multiset_eq_iff Nat.add_mult_distrib)  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
943  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
944  | 
lemma repeat_mset_distrib2[simp]:  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
945  | 
"repeat_mset n (A + B) = repeat_mset n A + repeat_mset n B"  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
946  | 
by (auto simp: multiset_eq_iff add_mult_distrib2)  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
947  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
948  | 
lemma repeat_mset_replicate_mset[simp]:  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
949  | 
  "repeat_mset n {#a#} = replicate_mset n a"
 | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
950  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
951  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
952  | 
lemma repeat_mset_distrib_add_mset[simp]:  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
953  | 
"repeat_mset n (add_mset a A) = replicate_mset n a + repeat_mset n A"  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
954  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
955  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
956  | 
lemma repeat_mset_empty[simp]: "repeat_mset n {#} = {#}"
 | 
| 73393 | 957  | 
by transfer simp  | 
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
958  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
959  | 
|
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
960  | 
subsubsection \<open>Simprocs\<close>  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
961  | 
|
| 
65031
 
52e2c99f3711
use the cancellation simprocs directly
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65029 
diff
changeset
 | 
962  | 
lemma repeat_mset_iterate_add: \<open>repeat_mset n M = iterate_add n M\<close>  | 
| 
 
52e2c99f3711
use the cancellation simprocs directly
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65029 
diff
changeset
 | 
963  | 
unfolding iterate_add_def by (induction n) auto  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
964  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
965  | 
lemma mset_subseteq_add_iff1:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
966  | 
"j \<le> (i::nat) \<Longrightarrow> (repeat_mset i u + m \<subseteq># repeat_mset j u + n) = (repeat_mset (i-j) u + m \<subseteq># n)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
967  | 
by (auto simp add: subseteq_mset_def nat_le_add_iff1)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
968  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
969  | 
lemma mset_subseteq_add_iff2:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
970  | 
"i \<le> (j::nat) \<Longrightarrow> (repeat_mset i u + m \<subseteq># repeat_mset j u + n) = (m \<subseteq># repeat_mset (j-i) u + n)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
971  | 
by (auto simp add: subseteq_mset_def nat_le_add_iff2)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
972  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
973  | 
lemma mset_subset_add_iff1:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
974  | 
"j \<le> (i::nat) \<Longrightarrow> (repeat_mset i u + m \<subset># repeat_mset j u + n) = (repeat_mset (i-j) u + m \<subset># n)"  | 
| 
65031
 
52e2c99f3711
use the cancellation simprocs directly
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65029 
diff
changeset
 | 
975  | 
unfolding subset_mset_def repeat_mset_iterate_add  | 
| 
 
52e2c99f3711
use the cancellation simprocs directly
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65029 
diff
changeset
 | 
976  | 
by (simp add: iterate_add_eq_add_iff1 mset_subseteq_add_iff1[unfolded repeat_mset_iterate_add])  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
977  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
978  | 
lemma mset_subset_add_iff2:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
979  | 
"i \<le> (j::nat) \<Longrightarrow> (repeat_mset i u + m \<subset># repeat_mset j u + n) = (m \<subset># repeat_mset (j-i) u + n)"  | 
| 
65031
 
52e2c99f3711
use the cancellation simprocs directly
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65029 
diff
changeset
 | 
980  | 
unfolding subset_mset_def repeat_mset_iterate_add  | 
| 
 
52e2c99f3711
use the cancellation simprocs directly
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65029 
diff
changeset
 | 
981  | 
by (simp add: iterate_add_eq_add_iff2 mset_subseteq_add_iff2[unfolded repeat_mset_iterate_add])  | 
| 
65029
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
982  | 
|
| 69605 | 983  | 
ML_file \<open>multiset_simprocs.ML\<close>  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
984  | 
|
| 
65029
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
985  | 
lemma add_mset_replicate_mset_safe[cancelation_simproc_pre]: \<open>NO_MATCH {#} M \<Longrightarrow> add_mset a M = {#a#} + M\<close>
 | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
986  | 
by simp  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
987  | 
|
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
988  | 
declare repeat_mset_iterate_add[cancelation_simproc_pre]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
989  | 
|
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
990  | 
declare iterate_add_distrib[cancelation_simproc_pre]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
991  | 
declare repeat_mset_iterate_add[symmetric, cancelation_simproc_post]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
992  | 
|
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
993  | 
declare add_mset_not_empty[cancelation_simproc_eq_elim]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
994  | 
empty_not_add_mset[cancelation_simproc_eq_elim]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
995  | 
subset_mset.le_zero_eq[cancelation_simproc_eq_elim]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
996  | 
empty_not_add_mset[cancelation_simproc_eq_elim]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
997  | 
add_mset_not_empty[cancelation_simproc_eq_elim]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
998  | 
subset_mset.le_zero_eq[cancelation_simproc_eq_elim]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
999  | 
le_zero_eq[cancelation_simproc_eq_elim]  | 
| 
 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
65027 
diff
changeset
 | 
1000  | 
|
| 
65027
 
2b8583507891
renaming multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
64911 
diff
changeset
 | 
1001  | 
simproc_setup mseteq_cancel  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1002  | 
  ("(l::'a multiset) + m = n" | "(l::'a multiset) = m + n" |
 | 
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1003  | 
"add_mset a m = n" | "m = add_mset a n" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1004  | 
"replicate_mset p a = n" | "m = replicate_mset p a" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1005  | 
"repeat_mset p m = n" | "m = repeat_mset p m") =  | 
| 
78099
 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 
wenzelm 
parents: 
77987 
diff
changeset
 | 
1006  | 
\<open>K Cancel_Simprocs.eq_cancel\<close>  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1007  | 
|
| 
65027
 
2b8583507891
renaming multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
64911 
diff
changeset
 | 
1008  | 
simproc_setup msetsubset_cancel  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1009  | 
  ("(l::'a multiset) + m \<subset># n" | "(l::'a multiset) \<subset># m + n" |
 | 
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1010  | 
"add_mset a m \<subset># n" | "m \<subset># add_mset a n" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1011  | 
"replicate_mset p r \<subset># n" | "m \<subset># replicate_mset p r" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1012  | 
"repeat_mset p m \<subset># n" | "m \<subset># repeat_mset p m") =  | 
| 
78099
 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 
wenzelm 
parents: 
77987 
diff
changeset
 | 
1013  | 
\<open>K Multiset_Simprocs.subset_cancel_msets\<close>  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1014  | 
|
| 
65027
 
2b8583507891
renaming multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
64911 
diff
changeset
 | 
1015  | 
simproc_setup msetsubset_eq_cancel  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1016  | 
  ("(l::'a multiset) + m \<subseteq># n" | "(l::'a multiset) \<subseteq># m + n" |
 | 
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1017  | 
"add_mset a m \<subseteq># n" | "m \<subseteq># add_mset a n" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1018  | 
"replicate_mset p r \<subseteq># n" | "m \<subseteq># replicate_mset p r" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1019  | 
"repeat_mset p m \<subseteq># n" | "m \<subseteq># repeat_mset p m") =  | 
| 
78099
 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 
wenzelm 
parents: 
77987 
diff
changeset
 | 
1020  | 
\<open>K Multiset_Simprocs.subseteq_cancel_msets\<close>  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1021  | 
|
| 
65027
 
2b8583507891
renaming multiset simprocs
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
64911 
diff
changeset
 | 
1022  | 
simproc_setup msetdiff_cancel  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1023  | 
  ("((l::'a multiset) + m) - n" | "(l::'a multiset) - (m + n)" |
 | 
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1024  | 
"add_mset a m - n" | "m - add_mset a n" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1025  | 
"replicate_mset p r - n" | "m - replicate_mset p r" |  | 
| 
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
1026  | 
"repeat_mset p m - n" | "m - repeat_mset p m") =  | 
| 
78099
 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 
wenzelm 
parents: 
77987 
diff
changeset
 | 
1027  | 
\<open>K Cancel_Simprocs.diff_cancel\<close>  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1028  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1029  | 
|
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1030  | 
subsubsection \<open>Conditionally complete lattice\<close>  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1031  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1032  | 
instantiation multiset :: (type) Inf  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1033  | 
begin  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1034  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1035  | 
lift_definition Inf_multiset :: "'a multiset set \<Rightarrow> 'a multiset" is  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1036  | 
  "\<lambda>A i. if A = {} then 0 else Inf ((\<lambda>f. f i) ` A)"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1037  | 
proof -  | 
| 73270 | 1038  | 
  fix A :: "('a \<Rightarrow> nat) set"
 | 
1039  | 
  assume *: "\<And>f. f \<in> A \<Longrightarrow> finite {x. 0 < f x}"
 | 
|
1040  | 
  show \<open>finite {i. 0 < (if A = {} then 0 else INF f\<in>A. f i)}\<close>
 | 
|
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1041  | 
  proof (cases "A = {}")
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1042  | 
case False  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1043  | 
then obtain f where "f \<in> A" by blast  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1044  | 
    hence "{i. Inf ((\<lambda>f. f i) ` A) > 0} \<subseteq> {i. f i > 0}"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1045  | 
by (auto intro: less_le_trans[OF _ cInf_lower])  | 
| 73270 | 1046  | 
moreover from \<open>f \<in> A\<close> * have "finite \<dots>" by simp  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1047  | 
    ultimately have "finite {i. Inf ((\<lambda>f. f i) ` A) > 0}" by (rule finite_subset)
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1048  | 
with False show ?thesis by simp  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1049  | 
qed simp_all  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1050  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1051  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1052  | 
instance ..  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1053  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1054  | 
end  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1055  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1056  | 
lemma Inf_multiset_empty: "Inf {} = {#}"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1057  | 
by transfer simp_all  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1058  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1059  | 
lemma count_Inf_multiset_nonempty: "A \<noteq> {} \<Longrightarrow> count (Inf A) x = Inf ((\<lambda>X. count X x) ` A)"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1060  | 
by transfer simp_all  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1061  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1062  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1063  | 
instantiation multiset :: (type) Sup  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1064  | 
begin  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1065  | 
|
| 63360 | 1066  | 
definition Sup_multiset :: "'a multiset set \<Rightarrow> 'a multiset" where  | 
1067  | 
  "Sup_multiset A = (if A \<noteq> {} \<and> subset_mset.bdd_above A then
 | 
|
1068  | 
           Abs_multiset (\<lambda>i. Sup ((\<lambda>X. count X i) ` A)) else {#})"
 | 
|
1069  | 
||
1070  | 
lemma Sup_multiset_empty: "Sup {} = {#}"
 | 
|
1071  | 
by (simp add: Sup_multiset_def)  | 
|
1072  | 
||
| 73451 | 1073  | 
lemma Sup_multiset_unbounded: "\<not> subset_mset.bdd_above A \<Longrightarrow> Sup A = {#}"
 | 
| 63360 | 1074  | 
by (simp add: Sup_multiset_def)  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1075  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1076  | 
instance ..  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1077  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1078  | 
end  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1079  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1080  | 
lemma bdd_above_multiset_imp_bdd_above_count:  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1081  | 
assumes "subset_mset.bdd_above (A :: 'a multiset set)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1082  | 
shows "bdd_above ((\<lambda>X. count X x) ` A)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1083  | 
proof -  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1084  | 
from assms obtain Y where Y: "\<forall>X\<in>A. X \<subseteq># Y"  | 
| 73451 | 1085  | 
by (meson subset_mset.bdd_above.E)  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1086  | 
hence "count X x \<le> count Y x" if "X \<in> A" for X  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1087  | 
using that by (auto intro: mset_subset_eq_count)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1088  | 
thus ?thesis by (intro bdd_aboveI[of _ "count Y x"]) auto  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1089  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1090  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1091  | 
lemma bdd_above_multiset_imp_finite_support:  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1092  | 
  assumes "A \<noteq> {}" "subset_mset.bdd_above (A :: 'a multiset set)"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1093  | 
  shows   "finite (\<Union>X\<in>A. {x. count X x > 0})"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1094  | 
proof -  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1095  | 
from assms obtain Y where Y: "\<forall>X\<in>A. X \<subseteq># Y"  | 
| 73451 | 1096  | 
by (meson subset_mset.bdd_above.E)  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1097  | 
hence "count X x \<le> count Y x" if "X \<in> A" for X x  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1098  | 
using that by (auto intro: mset_subset_eq_count)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1099  | 
  hence "(\<Union>X\<in>A. {x. count X x > 0}) \<subseteq> {x. count Y x > 0}"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1100  | 
by safe (erule less_le_trans)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1101  | 
moreover have "finite \<dots>" by simp  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1102  | 
ultimately show ?thesis by (rule finite_subset)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1103  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1104  | 
|
| 63360 | 1105  | 
lemma Sup_multiset_in_multiset:  | 
| 73270 | 1106  | 
  \<open>finite {i. 0 < (SUP M\<in>A. count M i)}\<close>
 | 
1107  | 
  if \<open>A \<noteq> {}\<close> \<open>subset_mset.bdd_above A\<close>
 | 
|
1108  | 
proof -  | 
|
| 63360 | 1109  | 
  have "{i. Sup ((\<lambda>X. count X i) ` A) > 0} \<subseteq> (\<Union>X\<in>A. {i. 0 < count X i})"
 | 
1110  | 
proof safe  | 
|
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69107 
diff
changeset
 | 
1111  | 
fix i assume pos: "(SUP X\<in>A. count X i) > 0"  | 
| 63360 | 1112  | 
    show "i \<in> (\<Union>X\<in>A. {i. 0 < count X i})"
 | 
1113  | 
proof (rule ccontr)  | 
|
1114  | 
      assume "i \<notin> (\<Union>X\<in>A. {i. 0 < count X i})"
 | 
|
1115  | 
hence "\<forall>X\<in>A. count X i \<le> 0" by (auto simp: count_eq_zero_iff)  | 
|
| 73270 | 1116  | 
with that have "(SUP X\<in>A. count X i) \<le> 0"  | 
| 63360 | 1117  | 
by (intro cSup_least bdd_above_multiset_imp_bdd_above_count) auto  | 
1118  | 
with pos show False by simp  | 
|
1119  | 
qed  | 
|
1120  | 
qed  | 
|
| 73270 | 1121  | 
moreover from that have "finite \<dots>"  | 
1122  | 
by (rule bdd_above_multiset_imp_finite_support)  | 
|
1123  | 
  ultimately show "finite {i. Sup ((\<lambda>X. count X i) ` A) > 0}"
 | 
|
1124  | 
by (rule finite_subset)  | 
|
| 63360 | 1125  | 
qed  | 
1126  | 
||
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1127  | 
lemma count_Sup_multiset_nonempty:  | 
| 73270 | 1128  | 
\<open>count (Sup A) x = (SUP X\<in>A. count X x)\<close>  | 
1129  | 
  if \<open>A \<noteq> {}\<close> \<open>subset_mset.bdd_above A\<close>
 | 
|
1130  | 
using that by (simp add: Sup_multiset_def Sup_multiset_in_multiset count_Abs_multiset)  | 
|
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1131  | 
|
| 67398 | 1132  | 
interpretation subset_mset: conditionally_complete_lattice Inf Sup "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)" "(\<union>#)"  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1133  | 
proof  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1134  | 
fix X :: "'a multiset" and A  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1135  | 
assume "X \<in> A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1136  | 
show "Inf A \<subseteq># X"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1137  | 
proof (rule mset_subset_eqI)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1138  | 
fix x  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1139  | 
    from \<open>X \<in> A\<close> have "A \<noteq> {}" by auto
 | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69107 
diff
changeset
 | 
1140  | 
hence "count (Inf A) x = (INF X\<in>A. count X x)"  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1141  | 
by (simp add: count_Inf_multiset_nonempty)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1142  | 
also from \<open>X \<in> A\<close> have "\<dots> \<le> count X x"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1143  | 
by (intro cInf_lower) simp_all  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1144  | 
finally show "count (Inf A) x \<le> count X x" .  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1145  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1146  | 
next  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1147  | 
fix X :: "'a multiset" and A  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1148  | 
  assume nonempty: "A \<noteq> {}" and le: "\<And>Y. Y \<in> A \<Longrightarrow> X \<subseteq># Y"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1149  | 
show "X \<subseteq># Inf A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1150  | 
proof (rule mset_subset_eqI)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1151  | 
fix x  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69107 
diff
changeset
 | 
1152  | 
from nonempty have "count X x \<le> (INF X\<in>A. count X x)"  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1153  | 
by (intro cInf_greatest) (auto intro: mset_subset_eq_count le)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1154  | 
also from nonempty have "\<dots> = count (Inf A) x" by (simp add: count_Inf_multiset_nonempty)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1155  | 
finally show "count X x \<le> count (Inf A) x" .  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1156  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1157  | 
next  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1158  | 
fix X :: "'a multiset" and A  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1159  | 
assume X: "X \<in> A" and bdd: "subset_mset.bdd_above A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1160  | 
show "X \<subseteq># Sup A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1161  | 
proof (rule mset_subset_eqI)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1162  | 
fix x  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1163  | 
    from X have "A \<noteq> {}" by auto
 | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69107 
diff
changeset
 | 
1164  | 
have "count X x \<le> (SUP X\<in>A. count X x)"  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1165  | 
by (intro cSUP_upper X bdd_above_multiset_imp_bdd_above_count bdd)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1166  | 
    also from count_Sup_multiset_nonempty[OF \<open>A \<noteq> {}\<close> bdd]
 | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69107 
diff
changeset
 | 
1167  | 
have "(SUP X\<in>A. count X x) = count (Sup A) x" by simp  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1168  | 
finally show "count X x \<le> count (Sup A) x" .  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1169  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1170  | 
next  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1171  | 
fix X :: "'a multiset" and A  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1172  | 
  assume nonempty: "A \<noteq> {}" and ge: "\<And>Y. Y \<in> A \<Longrightarrow> Y \<subseteq># X"
 | 
| 73451 | 1173  | 
from ge have bdd: "subset_mset.bdd_above A"  | 
1174  | 
by blast  | 
|
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1175  | 
show "Sup A \<subseteq># X"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1176  | 
proof (rule mset_subset_eqI)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1177  | 
fix x  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1178  | 
    from count_Sup_multiset_nonempty[OF \<open>A \<noteq> {}\<close> bdd]
 | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69107 
diff
changeset
 | 
1179  | 
have "count (Sup A) x = (SUP X\<in>A. count X x)" .  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1180  | 
also from nonempty have "\<dots> \<le> count X x"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1181  | 
by (intro cSup_least) (auto intro: mset_subset_eq_count ge)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1182  | 
finally show "count (Sup A) x \<le> count X x" .  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1183  | 
qed  | 
| 
64585
 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 
haftmann 
parents: 
64531 
diff
changeset
 | 
1184  | 
qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1185  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1186  | 
lemma set_mset_Inf:  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1187  | 
  assumes "A \<noteq> {}"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1188  | 
shows "set_mset (Inf A) = (\<Inter>X\<in>A. set_mset X)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1189  | 
proof safe  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1190  | 
fix x X assume "x \<in># Inf A" "X \<in> A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1191  | 
  hence nonempty: "A \<noteq> {}" by (auto simp: Inf_multiset_empty)
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1192  | 
  from \<open>x \<in># Inf A\<close> have "{#x#} \<subseteq># Inf A" by auto
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1193  | 
also from \<open>X \<in> A\<close> have "\<dots> \<subseteq># X" by (rule subset_mset.cInf_lower) simp_all  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1194  | 
finally show "x \<in># X" by simp  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1195  | 
next  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1196  | 
fix x assume x: "x \<in> (\<Inter>X\<in>A. set_mset X)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1197  | 
  hence "{#x#} \<subseteq># X" if "X \<in> A" for X using that by auto
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1198  | 
  from assms and this have "{#x#} \<subseteq># Inf A" by (rule subset_mset.cInf_greatest)
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1199  | 
thus "x \<in># Inf A" by simp  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1200  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1201  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1202  | 
lemma in_Inf_multiset_iff:  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1203  | 
  assumes "A \<noteq> {}"
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1204  | 
shows "x \<in># Inf A \<longleftrightarrow> (\<forall>X\<in>A. x \<in># X)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1205  | 
proof -  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1206  | 
from assms have "set_mset (Inf A) = (\<Inter>X\<in>A. set_mset X)" by (rule set_mset_Inf)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1207  | 
also have "x \<in> \<dots> \<longleftrightarrow> (\<forall>X\<in>A. x \<in># X)" by simp  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1208  | 
finally show ?thesis .  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1209  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1210  | 
|
| 63360 | 1211  | 
lemma in_Inf_multisetD: "x \<in># Inf A \<Longrightarrow> X \<in> A \<Longrightarrow> x \<in># X"  | 
1212  | 
by (subst (asm) in_Inf_multiset_iff) auto  | 
|
1213  | 
||
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1214  | 
lemma set_mset_Sup:  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1215  | 
assumes "subset_mset.bdd_above A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1216  | 
shows "set_mset (Sup A) = (\<Union>X\<in>A. set_mset X)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1217  | 
proof safe  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1218  | 
fix x assume "x \<in># Sup A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1219  | 
  hence nonempty: "A \<noteq> {}" by (auto simp: Sup_multiset_empty)
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1220  | 
show "x \<in> (\<Union>X\<in>A. set_mset X)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1221  | 
proof (rule ccontr)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1222  | 
assume x: "x \<notin> (\<Union>X\<in>A. set_mset X)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1223  | 
have "count X x \<le> count (Sup A) x" if "X \<in> A" for X x  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1224  | 
using that by (intro mset_subset_eq_count subset_mset.cSup_upper assms)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1225  | 
    with x have "X \<subseteq># Sup A - {#x#}" if "X \<in> A" for X
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1226  | 
using that by (auto simp: subseteq_mset_def algebra_simps not_in_iff)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1227  | 
    hence "Sup A \<subseteq># Sup A - {#x#}" by (intro subset_mset.cSup_least nonempty)
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1228  | 
with \<open>x \<in># Sup A\<close> show False  | 
| 68406 | 1229  | 
by (auto simp: subseteq_mset_def simp flip: count_greater_zero_iff  | 
1230  | 
dest!: spec[of _ x])  | 
|
| 
63358
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1231  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1232  | 
next  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1233  | 
fix x X assume "x \<in> set_mset X" "X \<in> A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1234  | 
  hence "{#x#} \<subseteq># X" by auto
 | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1235  | 
also have "X \<subseteq># Sup A" by (intro subset_mset.cSup_upper \<open>X \<in> A\<close> assms)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1236  | 
finally show "x \<in> set_mset (Sup A)" by simp  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1237  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1238  | 
|
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1239  | 
lemma in_Sup_multiset_iff:  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1240  | 
assumes "subset_mset.bdd_above A"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1241  | 
shows "x \<in># Sup A \<longleftrightarrow> (\<exists>X\<in>A. x \<in># X)"  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1242  | 
proof -  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1243  | 
from assms have "set_mset (Sup A) = (\<Union>X\<in>A. set_mset X)" by (rule set_mset_Sup)  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1244  | 
also have "x \<in> \<dots> \<longleftrightarrow> (\<exists>X\<in>A. x \<in># X)" by simp  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1245  | 
finally show ?thesis .  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1246  | 
qed  | 
| 
 
a500677d4cec
Conditionally complete lattice of multisets
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63310 
diff
changeset
 | 
1247  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1248  | 
lemma in_Sup_multisetD:  | 
| 63360 | 1249  | 
assumes "x \<in># Sup A"  | 
1250  | 
shows "\<exists>X\<in>A. x \<in># X"  | 
|
1251  | 
proof -  | 
|
1252  | 
have "subset_mset.bdd_above A"  | 
|
1253  | 
by (rule ccontr) (insert assms, simp_all add: Sup_multiset_unbounded)  | 
|
1254  | 
with assms show ?thesis by (simp add: in_Sup_multiset_iff)  | 
|
| 
63534
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
63524 
diff
changeset
 | 
1255  | 
qed  | 
| 
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
63524 
diff
changeset
 | 
1256  | 
|
| 67398 | 1257  | 
interpretation subset_mset: distrib_lattice "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)" "(\<union>#)"  | 
| 
63534
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
63524 
diff
changeset
 | 
1258  | 
proof  | 
| 
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
63524 
diff
changeset
 | 
1259  | 
fix A B C :: "'a multiset"  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
1260  | 
show "A \<union># (B \<inter># C) = A \<union># B \<inter># (A \<union># C)"  | 
| 
63534
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
63524 
diff
changeset
 | 
1261  | 
by (intro multiset_eqI) simp_all  | 
| 
64585
 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 
haftmann 
parents: 
64531 
diff
changeset
 | 
1262  | 
qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close>  | 
| 63360 | 1263  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1264  | 
|
| 60500 | 1265  | 
subsubsection \<open>Filter (with comprehension syntax)\<close>  | 
1266  | 
||
1267  | 
text \<open>Multiset comprehension\<close>  | 
|
| 
41069
 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 
haftmann 
parents: 
40968 
diff
changeset
 | 
1268  | 
|
| 
59998
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1269  | 
lift_definition filter_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"
 | 
| 
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1270  | 
is "\<lambda>P M. \<lambda>x. if P x then M x else 0"  | 
| 
47429
 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 
bulwahn 
parents: 
47308 
diff
changeset
 | 
1271  | 
by (rule filter_preserves_multiset)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1272  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1273  | 
syntax (ASCII)  | 
| 63689 | 1274  | 
  "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{#_ :# _./ _#})")
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1275  | 
syntax  | 
| 63689 | 1276  | 
  "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{#_ \<in># _./ _#})")
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1277  | 
translations  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1278  | 
  "{#x \<in># M. P#}" == "CONST filter_mset (\<lambda>x. P) M"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1279  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1280  | 
lemma count_filter_mset [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1281  | 
"count (filter_mset P M) a = (if P a then count M a else 0)"  | 
| 
59998
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1282  | 
by (simp add: filter_mset.rep_eq)  | 
| 
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1283  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1284  | 
lemma set_mset_filter [simp]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1285  | 
  "set_mset (filter_mset P M) = {a \<in> set_mset M. P a}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1286  | 
by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_filter_mset) simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1287  | 
|
| 60606 | 1288  | 
lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}"
 | 
| 
59998
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1289  | 
by (rule multiset_eqI) simp  | 
| 
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1290  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1291  | 
lemma filter_single_mset: "filter_mset P {#x#} = (if P x then {#x#} else {#})"
 | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
1292  | 
by (rule multiset_eqI) simp  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1293  | 
|
| 60606 | 1294  | 
lemma filter_union_mset [simp]: "filter_mset P (M + N) = filter_mset P M + filter_mset P N"  | 
| 
41069
 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 
haftmann 
parents: 
40968 
diff
changeset
 | 
1295  | 
by (rule multiset_eqI) simp  | 
| 
 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 
haftmann 
parents: 
40968 
diff
changeset
 | 
1296  | 
|
| 60606 | 1297  | 
lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N"  | 
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
1298  | 
by (rule multiset_eqI) simp  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
1299  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
1300  | 
lemma filter_inter_mset [simp]: "filter_mset P (M \<inter># N) = filter_mset P M \<inter># filter_mset P N"  | 
| 
41069
 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 
haftmann 
parents: 
40968 
diff
changeset
 | 
1301  | 
by (rule multiset_eqI) simp  | 
| 
 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 
haftmann 
parents: 
40968 
diff
changeset
 | 
1302  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
1303  | 
lemma filter_sup_mset[simp]: "filter_mset P (A \<union># B) = filter_mset P A \<union># filter_mset P B"  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
1304  | 
by (rule multiset_eqI) simp  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
1305  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1306  | 
lemma filter_mset_add_mset [simp]:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1307  | 
"filter_mset P (add_mset x A) =  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
1308  | 
(if P x then add_mset x (filter_mset P A) else filter_mset P A)"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1309  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1310  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1311  | 
lemma multiset_filter_subset[simp]: "filter_mset f M \<subseteq># M"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1312  | 
by (simp add: mset_subset_eqI)  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
1313  | 
|
| 60606 | 1314  | 
lemma multiset_filter_mono:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1315  | 
assumes "A \<subseteq># B"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1316  | 
shows "filter_mset f A \<subseteq># filter_mset f B"  | 
| 58035 | 1317  | 
proof -  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1318  | 
from assms[unfolded mset_subset_eq_exists_conv]  | 
| 58035 | 1319  | 
obtain C where B: "B = A + C" by auto  | 
1320  | 
show ?thesis unfolding B by auto  | 
|
1321  | 
qed  | 
|
1322  | 
||
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1323  | 
lemma filter_mset_eq_conv:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1324  | 
"filter_mset P M = N \<longleftrightarrow> N \<subseteq># M \<and> (\<forall>b\<in>#N. P b) \<and> (\<forall>a\<in>#M - N. \<not> P a)" (is "?P \<longleftrightarrow> ?Q")  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1325  | 
proof  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1326  | 
assume ?P then show ?Q by auto (simp add: multiset_eq_iff in_diff_count)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1327  | 
next  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1328  | 
assume ?Q  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1329  | 
then obtain Q where M: "M = N + Q"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1330  | 
by (auto simp add: mset_subset_eq_exists_conv)  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1331  | 
then have MN: "M - N = Q" by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1332  | 
show ?P  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1333  | 
proof (rule multiset_eqI)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1334  | 
fix a  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1335  | 
from \<open>?Q\<close> MN have *: "\<not> P a \<Longrightarrow> a \<notin># N" "P a \<Longrightarrow> a \<notin># Q"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1336  | 
by auto  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1337  | 
show "count (filter_mset P M) a = count N a"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1338  | 
proof (cases "a \<in># M")  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1339  | 
case True  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1340  | 
with * show ?thesis  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1341  | 
by (simp add: not_in_iff M)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1342  | 
next  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1343  | 
case False then have "count M a = 0"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1344  | 
by (simp add: not_in_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1345  | 
with M show ?thesis by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1346  | 
qed  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1347  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1348  | 
qed  | 
| 59813 | 1349  | 
|
| 64077 | 1350  | 
lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x \<in># M. Q x \<and> P x#}"
 | 
1351  | 
by (auto simp: multiset_eq_iff)  | 
|
1352  | 
||
| 64418 | 1353  | 
lemma  | 
1354  | 
  filter_mset_True[simp]: "{#y \<in># M. True#} = M" and
 | 
|
1355  | 
  filter_mset_False[simp]: "{#y \<in># M. False#} = {#}"
 | 
|
1356  | 
by (auto simp: multiset_eq_iff)  | 
|
1357  | 
||
| 75457 | 1358  | 
lemma filter_mset_cong0:  | 
1359  | 
assumes "\<And>x. x \<in># M \<Longrightarrow> f x \<longleftrightarrow> g x"  | 
|
1360  | 
shows "filter_mset f M = filter_mset g M"  | 
|
1361  | 
proof (rule subset_mset.antisym; unfold subseteq_mset_def; rule allI)  | 
|
1362  | 
fix x  | 
|
1363  | 
show "count (filter_mset f M) x \<le> count (filter_mset g M) x"  | 
|
1364  | 
using assms by (cases "x \<in># M") (simp_all add: not_in_iff)  | 
|
1365  | 
next  | 
|
1366  | 
fix x  | 
|
1367  | 
show "count (filter_mset g M) x \<le> count (filter_mset f M) x"  | 
|
1368  | 
using assms by (cases "x \<in># M") (simp_all add: not_in_iff)  | 
|
1369  | 
qed  | 
|
1370  | 
||
1371  | 
lemma filter_mset_cong:  | 
|
1372  | 
assumes "M = M'" and "\<And>x. x \<in># M' \<Longrightarrow> f x \<longleftrightarrow> g x"  | 
|
1373  | 
shows "filter_mset f M = filter_mset g M'"  | 
|
1374  | 
unfolding \<open>M = M'\<close>  | 
|
1375  | 
using assms by (auto intro: filter_mset_cong0)  | 
|
1376  | 
||
| 59813 | 1377  | 
|
| 60500 | 1378  | 
subsubsection \<open>Size\<close>  | 
| 10249 | 1379  | 
|
| 56656 | 1380  | 
definition wcount where "wcount f M = (\<lambda>x. count M x * Suc (f x))"  | 
1381  | 
||
1382  | 
lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a"  | 
|
1383  | 
by (auto simp: wcount_def add_mult_distrib)  | 
|
1384  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1385  | 
lemma wcount_add_mset:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1386  | 
"wcount f (add_mset x M) a = (if x = a then Suc (f a) else 0) + wcount f M a"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1387  | 
unfolding add_mset_add_single[of _ M] wcount_union by (auto simp: wcount_def)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1388  | 
|
| 56656 | 1389  | 
definition size_multiset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a multiset \<Rightarrow> nat" where
 | 
| 64267 | 1390  | 
"size_multiset f M = sum (wcount f M) (set_mset M)"  | 
| 56656 | 1391  | 
|
1392  | 
lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def]  | 
|
1393  | 
||
| 60606 | 1394  | 
instantiation multiset :: (type) size  | 
1395  | 
begin  | 
|
1396  | 
||
| 56656 | 1397  | 
definition size_multiset where  | 
1398  | 
size_multiset_overloaded_def: "size_multiset = Multiset.size_multiset (\<lambda>_. 0)"  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1399  | 
instance ..  | 
| 60606 | 1400  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1401  | 
end  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1402  | 
|
| 56656 | 1403  | 
lemmas size_multiset_overloaded_eq =  | 
1404  | 
size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified]  | 
|
1405  | 
||
1406  | 
lemma size_multiset_empty [simp]: "size_multiset f {#} = 0"
 | 
|
1407  | 
by (simp add: size_multiset_def)  | 
|
1408  | 
||
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
1409  | 
lemma size_empty [simp]: "size {#} = 0"
 | 
| 56656 | 1410  | 
by (simp add: size_multiset_overloaded_def)  | 
1411  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1412  | 
lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)"
 | 
| 56656 | 1413  | 
by (simp add: size_multiset_eq)  | 
| 10249 | 1414  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1415  | 
lemma size_single: "size {#b#} = 1"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1416  | 
by (simp add: size_multiset_overloaded_def size_multiset_single)  | 
| 56656 | 1417  | 
|
| 64267 | 1418  | 
lemma sum_wcount_Int:  | 
1419  | 
"finite A \<Longrightarrow> sum (wcount f N) (A \<inter> set_mset N) = sum (wcount f N) A"  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1420  | 
by (induct rule: finite_induct)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1421  | 
(simp_all add: Int_insert_left wcount_def count_eq_zero_iff)  | 
| 56656 | 1422  | 
|
1423  | 
lemma size_multiset_union [simp]:  | 
|
1424  | 
"size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N"  | 
|
| 64267 | 1425  | 
apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union)  | 
| 56656 | 1426  | 
apply (subst Int_commute)  | 
| 64267 | 1427  | 
apply (simp add: sum_wcount_Int)  | 
| 26178 | 1428  | 
done  | 
| 10249 | 1429  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1430  | 
lemma size_multiset_add_mset [simp]:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1431  | 
"size_multiset f (add_mset a M) = Suc (f a) + size_multiset f M"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1432  | 
unfolding add_mset_add_single[of _ M] size_multiset_union by (auto simp: size_multiset_single)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1433  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1434  | 
lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1435  | 
by (simp add: size_multiset_overloaded_def wcount_add_mset)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1436  | 
|
| 
28708
 
a1a436f09ec6
explicit check for pattern discipline before code translation
 
haftmann 
parents: 
28562 
diff
changeset
 | 
1437  | 
lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N"  | 
| 56656 | 1438  | 
by (auto simp add: size_multiset_overloaded_def)  | 
1439  | 
||
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1440  | 
lemma size_multiset_eq_0_iff_empty [iff]:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1441  | 
  "size_multiset f M = 0 \<longleftrightarrow> M = {#}"
 | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1442  | 
by (auto simp add: size_multiset_eq count_eq_zero_iff)  | 
| 10249 | 1443  | 
|
| 17161 | 1444  | 
lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
 | 
| 56656 | 1445  | 
by (auto simp add: size_multiset_overloaded_def)  | 
| 26016 | 1446  | 
|
1447  | 
lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
 | 
|
| 26178 | 1448  | 
by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty)  | 
| 10249 | 1449  | 
|
| 60607 | 1450  | 
lemma size_eq_Suc_imp_elem: "size M = Suc n \<Longrightarrow> \<exists>a. a \<in># M"  | 
| 56656 | 1451  | 
apply (unfold size_multiset_overloaded_eq)  | 
| 64267 | 1452  | 
apply (drule sum_SucD)  | 
| 26178 | 1453  | 
apply auto  | 
1454  | 
done  | 
|
| 10249 | 1455  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1456  | 
lemma size_eq_Suc_imp_eq_union:  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1457  | 
assumes "size M = Suc n"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1458  | 
shows "\<exists>a N. M = add_mset a N"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1459  | 
proof -  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1460  | 
from assms obtain a where "a \<in># M"  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1461  | 
by (erule size_eq_Suc_imp_elem [THEN exE])  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1462  | 
  then have "M = add_mset a (M - {#a#})" by simp
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1463  | 
then show ?thesis by blast  | 
| 23611 | 1464  | 
qed  | 
| 15869 | 1465  | 
|
| 60606 | 1466  | 
lemma size_mset_mono:  | 
1467  | 
fixes A B :: "'a multiset"  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1468  | 
assumes "A \<subseteq># B"  | 
| 60606 | 1469  | 
shows "size A \<le> size B"  | 
| 59949 | 1470  | 
proof -  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1471  | 
from assms[unfolded mset_subset_eq_exists_conv]  | 
| 59949 | 1472  | 
obtain C where B: "B = A + C" by auto  | 
| 60606 | 1473  | 
show ?thesis unfolding B by (induct C) auto  | 
| 59949 | 1474  | 
qed  | 
1475  | 
||
| 
59998
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1476  | 
lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \<le> size M"  | 
| 59949 | 1477  | 
by (rule size_mset_mono[OF multiset_filter_subset])  | 
1478  | 
||
1479  | 
lemma size_Diff_submset:  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1480  | 
"M \<subseteq># M' \<Longrightarrow> size (M' - M) = size M' - size(M::'a multiset)"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1481  | 
by (metis add_diff_cancel_left' size_union mset_subset_eq_exists_conv)  | 
| 26016 | 1482  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1483  | 
|
| 60500 | 1484  | 
subsection \<open>Induction and case splits\<close>  | 
| 10249 | 1485  | 
|
| 18258 | 1486  | 
theorem multiset_induct [case_names empty add, induct type: multiset]:  | 
| 48009 | 1487  | 
  assumes empty: "P {#}"
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1488  | 
assumes add: "\<And>x M. P M \<Longrightarrow> P (add_mset x M)"  | 
| 48009 | 1489  | 
shows "P M"  | 
| 65545 | 1490  | 
proof (induct "size M" arbitrary: M)  | 
| 48009 | 1491  | 
case 0 thus "P M" by (simp add: empty)  | 
1492  | 
next  | 
|
1493  | 
case (Suc k)  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1494  | 
obtain N x where "M = add_mset x N"  | 
| 60500 | 1495  | 
using \<open>Suc k = size M\<close> [symmetric]  | 
| 48009 | 1496  | 
using size_eq_Suc_imp_eq_union by fast  | 
1497  | 
with Suc add show "P M" by simp  | 
|
| 10249 | 1498  | 
qed  | 
1499  | 
||
| 65545 | 1500  | 
lemma multiset_induct_min[case_names empty add]:  | 
1501  | 
fixes M :: "'a::linorder multiset"  | 
|
1502  | 
assumes  | 
|
1503  | 
    empty: "P {#}" and
 | 
|
1504  | 
add: "\<And>x M. P M \<Longrightarrow> (\<forall>y \<in># M. y \<ge> x) \<Longrightarrow> P (add_mset x M)"  | 
|
1505  | 
shows "P M"  | 
|
1506  | 
proof (induct "size M" arbitrary: M)  | 
|
1507  | 
case (Suc k)  | 
|
1508  | 
note ih = this(1) and Sk_eq_sz_M = this(2)  | 
|
1509  | 
||
| 66425 | 1510  | 
let ?y = "Min_mset M"  | 
| 65545 | 1511  | 
  let ?N = "M - {#?y#}"
 | 
1512  | 
||
1513  | 
have M: "M = add_mset ?y ?N"  | 
|
1514  | 
by (metis Min_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero  | 
|
1515  | 
set_mset_eq_empty_iff size_empty)  | 
|
1516  | 
show ?case  | 
|
1517  | 
by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset,  | 
|
1518  | 
meson Min_le finite_set_mset in_diffD)  | 
|
1519  | 
qed (simp add: empty)  | 
|
1520  | 
||
1521  | 
lemma multiset_induct_max[case_names empty add]:  | 
|
1522  | 
fixes M :: "'a::linorder multiset"  | 
|
1523  | 
assumes  | 
|
1524  | 
    empty: "P {#}" and
 | 
|
1525  | 
add: "\<And>x M. P M \<Longrightarrow> (\<forall>y \<in># M. y \<le> x) \<Longrightarrow> P (add_mset x M)"  | 
|
1526  | 
shows "P M"  | 
|
1527  | 
proof (induct "size M" arbitrary: M)  | 
|
1528  | 
case (Suc k)  | 
|
1529  | 
note ih = this(1) and Sk_eq_sz_M = this(2)  | 
|
1530  | 
||
| 66425 | 1531  | 
let ?y = "Max_mset M"  | 
| 65545 | 1532  | 
  let ?N = "M - {#?y#}"
 | 
1533  | 
||
1534  | 
have M: "M = add_mset ?y ?N"  | 
|
1535  | 
by (metis Max_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero  | 
|
1536  | 
set_mset_eq_empty_iff size_empty)  | 
|
1537  | 
show ?case  | 
|
1538  | 
by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset,  | 
|
1539  | 
meson Max_ge finite_set_mset in_diffD)  | 
|
1540  | 
qed (simp add: empty)  | 
|
1541  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1542  | 
lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = add_mset a A"
 | 
| 26178 | 1543  | 
by (induct M) auto  | 
| 25610 | 1544  | 
|
| 55913 | 1545  | 
lemma multiset_cases [cases type]:  | 
1546  | 
  obtains (empty) "M = {#}"
 | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1547  | 
| (add) x N where "M = add_mset x N"  | 
| 63092 | 1548  | 
by (induct M) simp_all  | 
| 25610 | 1549  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1550  | 
lemma multi_drop_mem_not_eq: "c \<in># B \<Longrightarrow> B - {#c#} \<noteq> B"
 | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1551  | 
by (cases "B = {#}") (auto dest: multi_member_split)
 | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1552  | 
|
| 68992 | 1553  | 
lemma union_filter_mset_complement[simp]:  | 
1554  | 
"\<forall>x. P x = (\<not> Q x) \<Longrightarrow> filter_mset P M + filter_mset Q M = M"  | 
|
1555  | 
by (subst multiset_eq_iff) auto  | 
|
1556  | 
||
| 66494 | 1557  | 
lemma multiset_partition: "M = {#x \<in># M. P x#} + {#x \<in># M. \<not> P x#}"
 | 
| 68992 | 1558  | 
by simp  | 
| 66494 | 1559  | 
|
1560  | 
lemma mset_subset_size: "A \<subset># B \<Longrightarrow> size A < size B"  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1561  | 
proof (induct A arbitrary: B)  | 
| 66494 | 1562  | 
case empty  | 
1563  | 
then show ?case  | 
|
1564  | 
using nonempty_has_size by auto  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1565  | 
next  | 
| 66494 | 1566  | 
case (add x A)  | 
1567  | 
have "add_mset x A \<subseteq># B"  | 
|
1568  | 
by (meson add.prems subset_mset_def)  | 
|
1569  | 
then show ?case  | 
|
1570  | 
by (metis (no_types) add.prems add.right_neutral add_diff_cancel_left' leD nat_neq_iff  | 
|
1571  | 
size_Diff_submset size_eq_0_iff_empty size_mset_mono subset_mset.le_iff_add subset_mset_def)  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1572  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1573  | 
|
| 59949 | 1574  | 
lemma size_1_singleton_mset: "size M = 1 \<Longrightarrow> \<exists>a. M = {#a#}"
 | 
| 66494 | 1575  | 
by (cases M) auto  | 
| 59949 | 1576  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1577  | 
|
| 60500 | 1578  | 
subsubsection \<open>Strong induction and subset induction for multisets\<close>  | 
1579  | 
||
1580  | 
text \<open>Well-foundedness of strict subset relation\<close>  | 
|
| 58098 | 1581  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1582  | 
lemma wf_subset_mset_rel: "wf {(M, N :: 'a multiset). M \<subset># N}"
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1583  | 
apply (rule wf_measure [THEN wf_subset, where f1=size])  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1584  | 
apply (clarsimp simp: measure_def inv_image_def mset_subset_size)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1585  | 
done  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1586  | 
|
| 76300 | 1587  | 
lemma wfP_subset_mset[simp]: "wfP (\<subset>#)"  | 
1588  | 
by (rule wf_subset_mset_rel[to_pred])  | 
|
1589  | 
||
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1590  | 
lemma full_multiset_induct [case_names less]:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1591  | 
assumes ih: "\<And>B. \<forall>(A::'a multiset). A \<subset># B \<longrightarrow> P A \<Longrightarrow> P B"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1592  | 
shows "P B"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1593  | 
apply (rule wf_subset_mset_rel [THEN wf_induct])  | 
| 58098 | 1594  | 
apply (rule ih, auto)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1595  | 
done  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1596  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1597  | 
lemma multi_subset_induct [consumes 2, case_names empty add]:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1598  | 
assumes "F \<subseteq># A"  | 
| 60606 | 1599  | 
    and empty: "P {#}"
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1600  | 
and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (add_mset a F)"  | 
| 60606 | 1601  | 
shows "P F"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1602  | 
proof -  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1603  | 
from \<open>F \<subseteq># A\<close>  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1604  | 
show ?thesis  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1605  | 
proof (induct F)  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1606  | 
    show "P {#}" by fact
 | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1607  | 
next  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1608  | 
fix x F  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1609  | 
assume P: "F \<subseteq># A \<Longrightarrow> P F" and i: "add_mset x F \<subseteq># A"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1610  | 
show "P (add_mset x F)"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1611  | 
proof (rule insert)  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1612  | 
from i show "x \<in># A" by (auto dest: mset_subset_eq_insertD)  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
1613  | 
from i have "F \<subseteq># A" by (auto dest: mset_subset_eq_insertD)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1614  | 
with P show "P F" .  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1615  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1616  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1617  | 
qed  | 
| 26145 | 1618  | 
|
| 17161 | 1619  | 
|
| 
75467
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1620  | 
subsection \<open>Least and greatest elements\<close>  | 
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1621  | 
|
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1622  | 
context begin  | 
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1623  | 
|
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1624  | 
qualified lemma  | 
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1625  | 
assumes  | 
| 
77699
 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 
desharna 
parents: 
77688 
diff
changeset
 | 
1626  | 
    "M \<noteq> {#}" and
 | 
| 
76754
 
b5f4ae037fe2
used transp_on in assumptions of lemmas Multiset.bex_(least|greatest)_element
 
desharna 
parents: 
76749 
diff
changeset
 | 
1627  | 
"transp_on (set_mset M) R" and  | 
| 
77699
 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 
desharna 
parents: 
77688 
diff
changeset
 | 
1628  | 
"totalp_on (set_mset M) R"  | 
| 
75467
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1629  | 
shows  | 
| 
77699
 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 
desharna 
parents: 
77688 
diff
changeset
 | 
1630  | 
bex_least_element: "(\<exists>l \<in># M. \<forall>x \<in># M. x \<noteq> l \<longrightarrow> R l x)" and  | 
| 
 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 
desharna 
parents: 
77688 
diff
changeset
 | 
1631  | 
bex_greatest_element: "(\<exists>g \<in># M. \<forall>x \<in># M. x \<noteq> g \<longrightarrow> R x g)"  | 
| 
75467
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1632  | 
using assms  | 
| 
77699
 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 
desharna 
parents: 
77688 
diff
changeset
 | 
1633  | 
by (auto intro: Finite_Set.bex_least_element Finite_Set.bex_greatest_element)  | 
| 
75467
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1634  | 
|
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1635  | 
end  | 
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1636  | 
|
| 
 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 
desharna 
parents: 
75459 
diff
changeset
 | 
1637  | 
|
| 60500 | 1638  | 
subsection \<open>The fold combinator\<close>  | 
| 48023 | 1639  | 
|
| 
59998
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1640  | 
definition fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b"
 | 
| 48023 | 1641  | 
where  | 
| 60495 | 1642  | 
"fold_mset f s M = Finite_Set.fold (\<lambda>x. f x ^^ count M x) s (set_mset M)"  | 
| 48023 | 1643  | 
|
| 60606 | 1644  | 
lemma fold_mset_empty [simp]: "fold_mset f s {#} = s"
 | 
| 
59998
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
1645  | 
by (simp add: fold_mset_def)  | 
| 48023 | 1646  | 
|
1647  | 
context comp_fun_commute  | 
|
1648  | 
begin  | 
|
1649  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1650  | 
lemma fold_mset_add_mset [simp]: "fold_mset f s (add_mset x M) = f x (fold_mset f s M)"  | 
| 
49822
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1651  | 
proof -  | 
| 
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1652  | 
interpret mset: comp_fun_commute "\<lambda>y. f y ^^ count M y"  | 
| 
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1653  | 
by (fact comp_fun_commute_funpow)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1654  | 
interpret mset_union: comp_fun_commute "\<lambda>y. f y ^^ count (add_mset x M) y"  | 
| 
49822
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1655  | 
by (fact comp_fun_commute_funpow)  | 
| 
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1656  | 
show ?thesis  | 
| 60495 | 1657  | 
proof (cases "x \<in> set_mset M")  | 
| 
49822
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1658  | 
case False  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1659  | 
then have *: "count (add_mset x M) x = 1"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1660  | 
by (simp add: not_in_iff)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1661  | 
from False have "Finite_Set.fold (\<lambda>y. f y ^^ count (add_mset x M) y) s (set_mset M) =  | 
| 60495 | 1662  | 
Finite_Set.fold (\<lambda>y. f y ^^ count M y) s (set_mset M)"  | 
| 73832 | 1663  | 
by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow)  | 
| 
49822
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1664  | 
with False * show ?thesis  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1665  | 
by (simp add: fold_mset_def del: count_add_mset)  | 
| 48023 | 1666  | 
next  | 
| 
49822
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1667  | 
case True  | 
| 63040 | 1668  | 
    define N where "N = set_mset M - {x}"
 | 
| 60495 | 1669  | 
from N_def True have *: "set_mset M = insert x N" "x \<notin> N" "finite N" by auto  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1670  | 
then have "Finite_Set.fold (\<lambda>y. f y ^^ count (add_mset x M) y) s N =  | 
| 
49822
 
0cfc1651be25
simplified construction of fold combinator on multisets;
 
haftmann 
parents: 
49717 
diff
changeset
 | 
1671  | 
Finite_Set.fold (\<lambda>y. f y ^^ count M y) s N"  | 
| 73832 | 1672  | 
by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1673  | 
with * show ?thesis by (simp add: fold_mset_def del: count_add_mset) simp  | 
| 48023 | 1674  | 
qed  | 
1675  | 
qed  | 
|
1676  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1677  | 
corollary fold_mset_single: "fold_mset f s {#x#} = f x s"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1678  | 
by simp  | 
| 48023 | 1679  | 
|
| 60606 | 1680  | 
lemma fold_mset_fun_left_comm: "f x (fold_mset f s M) = fold_mset f (f x s) M"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1681  | 
by (induct M) (simp_all add: fun_left_comm)  | 
| 48023 | 1682  | 
|
| 60606 | 1683  | 
lemma fold_mset_union [simp]: "fold_mset f s (M + N) = fold_mset f (fold_mset f s M) N"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1684  | 
by (induct M) (simp_all add: fold_mset_fun_left_comm)  | 
| 48023 | 1685  | 
|
1686  | 
lemma fold_mset_fusion:  | 
|
1687  | 
assumes "comp_fun_commute g"  | 
|
| 60606 | 1688  | 
and *: "\<And>x y. h (g x y) = f x (h y)"  | 
1689  | 
shows "h (fold_mset g w A) = fold_mset f (h w) A"  | 
|
| 48023 | 1690  | 
proof -  | 
1691  | 
interpret comp_fun_commute g by (fact assms)  | 
|
| 60606 | 1692  | 
from * show ?thesis by (induct A) auto  | 
| 48023 | 1693  | 
qed  | 
1694  | 
||
1695  | 
end  | 
|
1696  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1697  | 
lemma union_fold_mset_add_mset: "A + B = fold_mset add_mset A B"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1698  | 
proof -  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1699  | 
interpret comp_fun_commute add_mset  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1700  | 
by standard auto  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1701  | 
show ?thesis  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1702  | 
by (induction B) auto  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1703  | 
qed  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1704  | 
|
| 60500 | 1705  | 
text \<open>  | 
| 48023 | 1706  | 
A note on code generation: When defining some function containing a  | 
| 69593 | 1707  | 
subterm \<^term>\<open>fold_mset F\<close>, code generation is not automatic. When  | 
| 61585 | 1708  | 
interpreting locale \<open>left_commutative\<close> with \<open>F\<close>, the  | 
| 69593 | 1709  | 
would be code thms for \<^const>\<open>fold_mset\<close> become thms like  | 
1710  | 
  \<^term>\<open>fold_mset F z {#} = z\<close> where \<open>F\<close> is not a pattern but
 | 
|
| 48023 | 1711  | 
contains defined symbols, i.e.\ is not a code thm. Hence a separate  | 
| 61585 | 1712  | 
constant with its own code thms needs to be introduced for \<open>F\<close>. See the image operator below.  | 
| 60500 | 1713  | 
\<close>  | 
1714  | 
||
1715  | 
||
1716  | 
subsection \<open>Image\<close>  | 
|
| 48023 | 1717  | 
|
1718  | 
definition image_mset :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
 | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1719  | 
  "image_mset f = fold_mset (add_mset \<circ> f) {#}"
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1720  | 
|
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1721  | 
lemma comp_fun_commute_mset_image: "comp_fun_commute (add_mset \<circ> f)"  | 
| 66494 | 1722  | 
by unfold_locales (simp add: fun_eq_iff)  | 
| 48023 | 1723  | 
|
1724  | 
lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
 | 
|
| 49823 | 1725  | 
by (simp add: image_mset_def)  | 
| 48023 | 1726  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1727  | 
lemma image_mset_single: "image_mset f {#x#} = {#f x#}"
 | 
| 66494 | 1728  | 
by (simp add: comp_fun_commute.fold_mset_add_mset comp_fun_commute_mset_image image_mset_def)  | 
| 48023 | 1729  | 
|
| 60606 | 1730  | 
lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N"  | 
| 49823 | 1731  | 
proof -  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1732  | 
interpret comp_fun_commute "add_mset \<circ> f"  | 
| 49823 | 1733  | 
by (fact comp_fun_commute_mset_image)  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
1734  | 
show ?thesis by (induct N) (simp_all add: image_mset_def)  | 
| 49823 | 1735  | 
qed  | 
1736  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1737  | 
corollary image_mset_add_mset [simp]:  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1738  | 
"image_mset f (add_mset a M) = add_mset (f a) (image_mset f M)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1739  | 
unfolding image_mset_union add_mset_add_single[of a M] by (simp add: image_mset_single)  | 
| 48023 | 1740  | 
|
| 60606 | 1741  | 
lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)"  | 
| 49823 | 1742  | 
by (induct M) simp_all  | 
| 48040 | 1743  | 
|
| 60606 | 1744  | 
lemma size_image_mset [simp]: "size (image_mset f M) = size M"  | 
| 49823 | 1745  | 
by (induct M) simp_all  | 
| 48023 | 1746  | 
|
| 60606 | 1747  | 
lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \<longleftrightarrow> M = {#}"
 | 
| 49823 | 1748  | 
by (cases M) auto  | 
| 48023 | 1749  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1750  | 
lemma image_mset_If:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1751  | 
"image_mset (\<lambda>x. if P x then f x else g x) A =  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1752  | 
image_mset f (filter_mset P A) + image_mset g (filter_mset (\<lambda>x. \<not>P x) A)"  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
1753  | 
by (induction A) auto  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1754  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1755  | 
lemma image_mset_Diff:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1756  | 
assumes "B \<subseteq># A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1757  | 
shows "image_mset f (A - B) = image_mset f A - image_mset f B"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1758  | 
proof -  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1759  | 
have "image_mset f (A - B + B) = image_mset f (A - B) + image_mset f B"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1760  | 
by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1761  | 
also from assms have "A - B + B = A"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1762  | 
by (simp add: subset_mset.diff_add)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1763  | 
finally show ?thesis by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1764  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1765  | 
|
| 73594 | 1766  | 
lemma count_image_mset:  | 
1767  | 
  \<open>count (image_mset f A) x = (\<Sum>y\<in>f -` {x} \<inter> set_mset A. count A y)\<close>
 | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1768  | 
proof (induction A)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1769  | 
case empty  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1770  | 
then show ?case by simp  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1771  | 
next  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1772  | 
case (add x A)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1773  | 
moreover have *: "(if x = y then Suc n else n) = n + (if x = y then 1 else 0)" for n y  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1774  | 
by simp  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1775  | 
ultimately show ?case  | 
| 66494 | 1776  | 
by (auto simp: sum.distrib intro!: sum.mono_neutral_left)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1777  | 
qed  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
1778  | 
|
| 73594 | 1779  | 
lemma count_image_mset':  | 
1780  | 
\<open>count (image_mset f X) y = (\<Sum>x | x \<in># X \<and> y = f x. count X x)\<close>  | 
|
1781  | 
by (auto simp add: count_image_mset simp flip: singleton_conv2 simp add: Collect_conj_eq ac_simps)  | 
|
1782  | 
||
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
1783  | 
lemma image_mset_subseteq_mono: "A \<subseteq># B \<Longrightarrow> image_mset f A \<subseteq># image_mset f B"  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
1784  | 
by (metis image_mset_union subset_mset.le_iff_add)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
1785  | 
|
| 65048 | 1786  | 
lemma image_mset_subset_mono: "M \<subset># N \<Longrightarrow> image_mset f M \<subset># image_mset f N"  | 
1787  | 
by (metis (no_types) Diff_eq_empty_iff_mset image_mset_Diff image_mset_is_empty_iff  | 
|
1788  | 
image_mset_subseteq_mono subset_mset.less_le_not_le)  | 
|
1789  | 
||
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1790  | 
syntax (ASCII)  | 
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1791  | 
  "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ :# _#})")
 | 
| 48023 | 1792  | 
syntax  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1793  | 
  "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ \<in># _#})")
 | 
| 59813 | 1794  | 
translations  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1795  | 
  "{#e. x \<in># M#}" \<rightleftharpoons> "CONST image_mset (\<lambda>x. e) M"
 | 
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1796  | 
|
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1797  | 
syntax (ASCII)  | 
| 
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1798  | 
  "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ :# _./ _#})")
 | 
| 48023 | 1799  | 
syntax  | 
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
1800  | 
  "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ \<in># _./ _#})")
 | 
| 59813 | 1801  | 
translations  | 
| 60606 | 1802  | 
  "{#e | x\<in>#M. P#}" \<rightharpoonup> "{#e. x \<in># {# x\<in>#M. P#}#}"
 | 
| 59813 | 1803  | 
|
| 60500 | 1804  | 
text \<open>  | 
| 69593 | 1805  | 
  This allows to write not just filters like \<^term>\<open>{#x\<in>#M. x<c#}\<close>
 | 
1806  | 
  but also images like \<^term>\<open>{#x+x. x\<in>#M #}\<close> and @{term [source]
 | 
|
| 60607 | 1807  | 
  "{#x+x|x\<in>#M. x<c#}"}, where the latter is currently displayed as
 | 
| 69593 | 1808  | 
  \<^term>\<open>{#x+x|x\<in>#M. x<c#}\<close>.
 | 
| 60500 | 1809  | 
\<close>  | 
| 48023 | 1810  | 
|
| 60495 | 1811  | 
lemma in_image_mset: "y \<in># {#f x. x \<in># M#} \<longleftrightarrow> y \<in> f ` set_mset M"
 | 
| 66494 | 1812  | 
by simp  | 
| 59813 | 1813  | 
|
| 
55467
 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 
blanchet 
parents: 
55417 
diff
changeset
 | 
1814  | 
functor image_mset: image_mset  | 
| 48023 | 1815  | 
proof -  | 
1816  | 
fix f g show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)"  | 
|
1817  | 
proof  | 
|
1818  | 
fix A  | 
|
1819  | 
show "(image_mset f \<circ> image_mset g) A = image_mset (f \<circ> g) A"  | 
|
1820  | 
by (induct A) simp_all  | 
|
1821  | 
qed  | 
|
1822  | 
show "image_mset id = id"  | 
|
1823  | 
proof  | 
|
1824  | 
fix A  | 
|
1825  | 
show "image_mset id A = id A"  | 
|
1826  | 
by (induct A) simp_all  | 
|
1827  | 
qed  | 
|
1828  | 
qed  | 
|
1829  | 
||
| 59813 | 1830  | 
declare  | 
1831  | 
image_mset.id [simp]  | 
|
1832  | 
image_mset.identity [simp]  | 
|
1833  | 
||
1834  | 
lemma image_mset_id[simp]: "image_mset id x = x"  | 
|
1835  | 
unfolding id_def by auto  | 
|
1836  | 
||
1837  | 
lemma image_mset_cong: "(\<And>x. x \<in># M \<Longrightarrow> f x = g x) \<Longrightarrow> {#f x. x \<in># M#} = {#g x. x \<in># M#}"
 | 
|
1838  | 
by (induct M) auto  | 
|
1839  | 
||
1840  | 
lemma image_mset_cong_pair:  | 
|
1841  | 
  "(\<forall>x y. (x, y) \<in># M \<longrightarrow> f x y = g x y) \<Longrightarrow> {#f x y. (x, y) \<in># M#} = {#g x y. (x, y) \<in># M#}"
 | 
|
1842  | 
by (metis image_mset_cong split_cong)  | 
|
| 49717 | 1843  | 
|
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64587 
diff
changeset
 | 
1844  | 
lemma image_mset_const_eq:  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64587 
diff
changeset
 | 
1845  | 
  "{#c. a \<in># M#} = replicate_mset (size M) c"
 | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64587 
diff
changeset
 | 
1846  | 
by (induct M) simp_all  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64587 
diff
changeset
 | 
1847  | 
|
| 75459 | 1848  | 
lemma image_mset_filter_mset_swap:  | 
1849  | 
"image_mset f (filter_mset (\<lambda>x. P (f x)) M) = filter_mset P (image_mset f M)"  | 
|
1850  | 
by (induction M rule: multiset_induct) simp_all  | 
|
1851  | 
||
| 48023 | 1852  | 
|
| 
75560
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1853  | 
lemma image_mset_eq_plusD:  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1854  | 
"image_mset f A = B + C \<Longrightarrow> \<exists>B' C'. A = B' + C' \<and> B = image_mset f B' \<and> C = image_mset f C'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1855  | 
proof (induction A arbitrary: B C)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1856  | 
case empty  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1857  | 
thus ?case by simp  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1858  | 
next  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1859  | 
case (add x A)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1860  | 
show ?case  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1861  | 
proof (cases "f x \<in># B")  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1862  | 
case True  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1863  | 
    with add.prems have "image_mset f A = (B - {#f x#}) + C"
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1864  | 
by (metis add_mset_remove_trivial image_mset_add_mset mset_subset_eq_single  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1865  | 
subset_mset.add_diff_assoc2)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1866  | 
thus ?thesis  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1867  | 
using add.IH add.prems by force  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1868  | 
next  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1869  | 
case False  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1870  | 
    with add.prems have "image_mset f A = B + (C - {#f x#})"
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1871  | 
by (metis diff_single_eq_union diff_union_single_conv image_mset_add_mset union_iff  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1872  | 
union_single_eq_member)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1873  | 
then show ?thesis  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1874  | 
using add.IH add.prems by force  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1875  | 
qed  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1876  | 
qed  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1877  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1878  | 
lemma image_mset_eq_image_mset_plusD:  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1879  | 
assumes "image_mset f A = image_mset f B + C" and inj_f: "inj_on f (set_mset A \<union> set_mset B)"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1880  | 
shows "\<exists>C'. A = B + C' \<and> C = image_mset f C'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1881  | 
using assms  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1882  | 
proof (induction A arbitrary: B C)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1883  | 
case empty  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1884  | 
thus ?case by simp  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1885  | 
next  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1886  | 
case (add x A)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1887  | 
show ?case  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1888  | 
proof (cases "x \<in># B")  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1889  | 
case True  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1890  | 
    with add.prems have "image_mset f A = image_mset f (B - {#x#}) + C"
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1891  | 
by (smt (verit, del_insts) add.left_commute add_cancel_right_left diff_union_cancelL  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1892  | 
diff_union_single_conv image_mset_union union_mset_add_mset_left  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1893  | 
union_mset_add_mset_right)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1894  | 
    with add.IH have "\<exists>M3'. A = B - {#x#} + M3' \<and> image_mset f M3' = C"
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1895  | 
by (smt (verit, del_insts) True Un_insert_left Un_insert_right add.prems(2) inj_on_insert  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1896  | 
insert_DiffM set_mset_add_mset_insert)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1897  | 
with True show ?thesis  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1898  | 
by auto  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1899  | 
next  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1900  | 
case False  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1901  | 
with add.prems(2) have "f x \<notin># image_mset f B"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1902  | 
by auto  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1903  | 
    with add.prems(1) have "image_mset f A = image_mset f B + (C - {#f x#})"
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1904  | 
by (metis (no_types, lifting) diff_union_single_conv image_eqI image_mset_Diff  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1905  | 
image_mset_single mset_subset_eq_single set_image_mset union_iff union_single_eq_diff  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1906  | 
union_single_eq_member)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1907  | 
    with add.prems(2) add.IH have "\<exists>M3'. A = B + M3' \<and> C - {#f x#} = image_mset f M3'"
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1908  | 
by auto  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1909  | 
then show ?thesis  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1910  | 
by (metis add.prems(1) add_diff_cancel_left' image_mset_Diff mset_subset_eq_add_left  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1911  | 
union_mset_add_mset_right)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1912  | 
qed  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1913  | 
qed  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1914  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1915  | 
lemma image_mset_eq_plus_image_msetD:  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1916  | 
"image_mset f A = B + image_mset f C \<Longrightarrow> inj_on f (set_mset A \<union> set_mset C) \<Longrightarrow>  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1917  | 
\<exists>B'. A = B' + C \<and> B = image_mset f B'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1918  | 
unfolding add.commute[of B] add.commute[of _ C]  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1919  | 
by (rule image_mset_eq_image_mset_plusD; assumption)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1920  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
1921  | 
|
| 60500 | 1922  | 
subsection \<open>Further conversions\<close>  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1923  | 
|
| 60515 | 1924  | 
primrec mset :: "'a list \<Rightarrow> 'a multiset" where  | 
1925  | 
  "mset [] = {#}" |
 | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
1926  | 
"mset (a # x) = add_mset a (mset x)"  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1927  | 
|
| 37107 | 1928  | 
lemma in_multiset_in_set:  | 
| 60515 | 1929  | 
"x \<in># mset xs \<longleftrightarrow> x \<in> set xs"  | 
| 37107 | 1930  | 
by (induct xs) simp_all  | 
1931  | 
||
| 60515 | 1932  | 
lemma count_mset:  | 
1933  | 
"count (mset xs) x = length (filter (\<lambda>y. x = y) xs)"  | 
|
| 37107 | 1934  | 
by (induct xs) simp_all  | 
1935  | 
||
| 60515 | 1936  | 
lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])"
 | 
| 59813 | 1937  | 
by (induct x) auto  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1938  | 
|
| 60515 | 1939  | 
lemma mset_zero_iff_right[simp]: "({#} = mset x) = (x = [])"
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1940  | 
by (induct x) auto  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1941  | 
|
| 
66276
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
1942  | 
lemma count_mset_gt_0: "x \<in> set xs \<Longrightarrow> count (mset xs) x > 0"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
1943  | 
by (induction xs) auto  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
1944  | 
|
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
1945  | 
lemma count_mset_0_iff [simp]: "count (mset xs) x = 0 \<longleftrightarrow> x \<notin> set xs"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
1946  | 
by (induction xs) auto  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
1947  | 
|
| 64077 | 1948  | 
lemma mset_single_iff[iff]: "mset xs = {#x#} \<longleftrightarrow> xs = [x]"
 | 
1949  | 
by (cases xs) auto  | 
|
1950  | 
||
1951  | 
lemma mset_single_iff_right[iff]: "{#x#} = mset xs \<longleftrightarrow> xs = [x]"
 | 
|
1952  | 
by (cases xs) auto  | 
|
1953  | 
||
| 64076 | 1954  | 
lemma set_mset_mset[simp]: "set_mset (mset xs) = set xs"  | 
1955  | 
by (induct xs) auto  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1956  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1957  | 
lemma set_mset_comp_mset [simp]: "set_mset \<circ> mset = set"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1958  | 
by (simp add: fun_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1959  | 
|
| 60515 | 1960  | 
lemma size_mset [simp]: "size (mset xs) = length xs"  | 
| 48012 | 1961  | 
by (induct xs) simp_all  | 
1962  | 
||
| 60606 | 1963  | 
lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys"  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
1964  | 
by (induct xs arbitrary: ys) auto  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1965  | 
|
| 68988 | 1966  | 
lemma mset_filter[simp]: "mset (filter P xs) = {#x \<in># mset xs. P x #}"
 | 
| 
40303
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
1967  | 
by (induct xs) simp_all  | 
| 
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
1968  | 
|
| 60515 | 1969  | 
lemma mset_rev [simp]:  | 
1970  | 
"mset (rev xs) = mset xs"  | 
|
| 40950 | 1971  | 
by (induct xs) simp_all  | 
1972  | 
||
| 60515 | 1973  | 
lemma surj_mset: "surj mset"  | 
| 76359 | 1974  | 
unfolding surj_def  | 
1975  | 
proof (rule allI)  | 
|
1976  | 
fix M  | 
|
1977  | 
show "\<exists>xs. M = mset xs"  | 
|
1978  | 
by (induction M) (auto intro: exI[of _ "_ # _"])  | 
|
1979  | 
qed  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1980  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
1981  | 
lemma distinct_count_atmost_1:  | 
| 60606 | 1982  | 
"distinct x = (\<forall>a. count (mset x) a = (if a \<in> set x then 1 else 0))"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1983  | 
proof (induct x)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1984  | 
case Nil then show ?case by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1985  | 
next  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1986  | 
case (Cons x xs) show ?case (is "?lhs \<longleftrightarrow> ?rhs")  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1987  | 
proof  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1988  | 
assume ?lhs then show ?rhs using Cons by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1989  | 
next  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1990  | 
assume ?rhs then have "x \<notin> set xs"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1991  | 
by (simp split: if_splits)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1992  | 
moreover from \<open>?rhs\<close> have "(\<forall>a. count (mset xs) a =  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1993  | 
(if a \<in> set xs then 1 else 0))"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1994  | 
by (auto split: if_splits simp add: count_eq_zero_iff)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1995  | 
ultimately show ?lhs using Cons by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1996  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1997  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1998  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
1999  | 
lemma mset_eq_setD:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2000  | 
assumes "mset xs = mset ys"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2001  | 
shows "set xs = set ys"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2002  | 
proof -  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2003  | 
from assms have "set_mset (mset xs) = set_mset (mset ys)"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2004  | 
by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2005  | 
then show ?thesis by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2006  | 
qed  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2007  | 
|
| 60515 | 2008  | 
lemma set_eq_iff_mset_eq_distinct:  | 
| 73301 | 2009  | 
\<open>distinct x \<Longrightarrow> distinct y \<Longrightarrow> set x = set y \<longleftrightarrow> mset x = mset y\<close>  | 
2010  | 
by (auto simp: multiset_eq_iff distinct_count_atmost_1)  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2011  | 
|
| 60515 | 2012  | 
lemma set_eq_iff_mset_remdups_eq:  | 
| 73301 | 2013  | 
\<open>set x = set y \<longleftrightarrow> mset (remdups x) = mset (remdups y)\<close>  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2014  | 
apply (rule iffI)  | 
| 60515 | 2015  | 
apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1])  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2016  | 
apply (drule distinct_remdups [THEN distinct_remdups  | 
| 60515 | 2017  | 
[THEN set_eq_iff_mset_eq_distinct [THEN iffD2]]])  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2018  | 
apply simp  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2019  | 
done  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2020  | 
|
| 73301 | 2021  | 
lemma mset_eq_imp_distinct_iff:  | 
2022  | 
\<open>distinct xs \<longleftrightarrow> distinct ys\<close> if \<open>mset xs = mset ys\<close>  | 
|
2023  | 
using that by (auto simp add: distinct_count_atmost_1 dest: mset_eq_setD)  | 
|
2024  | 
||
| 60607 | 2025  | 
lemma nth_mem_mset: "i < length ls \<Longrightarrow> (ls ! i) \<in># mset ls"  | 
| 60678 | 2026  | 
proof (induct ls arbitrary: i)  | 
2027  | 
case Nil  | 
|
2028  | 
then show ?case by simp  | 
|
2029  | 
next  | 
|
2030  | 
case Cons  | 
|
2031  | 
then show ?case by (cases i) auto  | 
|
2032  | 
qed  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2033  | 
|
| 60606 | 2034  | 
lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}"
 | 
| 60678 | 2035  | 
by (induct xs) (auto simp add: multiset_eq_iff)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2036  | 
|
| 60515 | 2037  | 
lemma mset_eq_length:  | 
2038  | 
assumes "mset xs = mset ys"  | 
|
| 37107 | 2039  | 
shows "length xs = length ys"  | 
| 60515 | 2040  | 
using assms by (metis size_mset)  | 
2041  | 
||
2042  | 
lemma mset_eq_length_filter:  | 
|
2043  | 
assumes "mset xs = mset ys"  | 
|
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2044  | 
shows "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) ys)"  | 
| 60515 | 2045  | 
using assms by (metis count_mset)  | 
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2046  | 
|
| 
45989
 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 
haftmann 
parents: 
45866 
diff
changeset
 | 
2047  | 
lemma fold_multiset_equiv:  | 
| 73706 | 2048  | 
\<open>List.fold f xs = List.fold f ys\<close>  | 
2049  | 
if f: \<open>\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x\<close>  | 
|
2050  | 
and \<open>mset xs = mset ys\<close>  | 
|
2051  | 
using f \<open>mset xs = mset ys\<close> [symmetric] proof (induction xs arbitrary: ys)  | 
|
| 60678 | 2052  | 
case Nil  | 
2053  | 
then show ?case by simp  | 
|
| 
45989
 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 
haftmann 
parents: 
45866 
diff
changeset
 | 
2054  | 
next  | 
| 
 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 
haftmann 
parents: 
45866 
diff
changeset
 | 
2055  | 
case (Cons x xs)  | 
| 73706 | 2056  | 
then have *: \<open>set ys = set (x # xs)\<close>  | 
| 60678 | 2057  | 
by (blast dest: mset_eq_setD)  | 
| 73706 | 2058  | 
have \<open>\<And>x y. x \<in> set ys \<Longrightarrow> y \<in> set ys \<Longrightarrow> f x \<circ> f y = f y \<circ> f x\<close>  | 
| 
45989
 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 
haftmann 
parents: 
45866 
diff
changeset
 | 
2059  | 
by (rule Cons.prems(1)) (simp_all add: *)  | 
| 73706 | 2060  | 
moreover from * have \<open>x \<in> set ys\<close>  | 
2061  | 
by simp  | 
|
2062  | 
ultimately have \<open>List.fold f ys = List.fold f (remove1 x ys) \<circ> f x\<close>  | 
|
2063  | 
by (fact fold_remove1_split)  | 
|
2064  | 
moreover from Cons.prems have \<open>List.fold f xs = List.fold f (remove1 x ys)\<close>  | 
|
2065  | 
by (auto intro: Cons.IH)  | 
|
2066  | 
ultimately show ?case  | 
|
| 60678 | 2067  | 
by simp  | 
| 73706 | 2068  | 
qed  | 
2069  | 
||
2070  | 
lemma fold_permuted_eq:  | 
|
2071  | 
\<open>List.fold (\<odot>) xs z = List.fold (\<odot>) ys z\<close>  | 
|
2072  | 
if \<open>mset xs = mset ys\<close>  | 
|
2073  | 
and \<open>P z\<close> and P: \<open>\<And>x z. x \<in> set xs \<Longrightarrow> P z \<Longrightarrow> P (x \<odot> z)\<close>  | 
|
2074  | 
and f: \<open>\<And>x y z. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> P z \<Longrightarrow> x \<odot> (y \<odot> z) = y \<odot> (x \<odot> z)\<close>  | 
|
2075  | 
for f (infixl \<open>\<odot>\<close> 70)  | 
|
2076  | 
using \<open>P z\<close> P f \<open>mset xs = mset ys\<close> [symmetric] proof (induction xs arbitrary: ys z)  | 
|
2077  | 
case Nil  | 
|
2078  | 
then show ?case by simp  | 
|
2079  | 
next  | 
|
2080  | 
case (Cons x xs)  | 
|
2081  | 
then have *: \<open>set ys = set (x # xs)\<close>  | 
|
2082  | 
by (blast dest: mset_eq_setD)  | 
|
2083  | 
have \<open>P z\<close>  | 
|
2084  | 
by (fact Cons.prems(1))  | 
|
2085  | 
moreover have \<open>\<And>x z. x \<in> set ys \<Longrightarrow> P z \<Longrightarrow> P (x \<odot> z)\<close>  | 
|
2086  | 
by (rule Cons.prems(2)) (simp_all add: *)  | 
|
2087  | 
moreover have \<open>\<And>x y z. x \<in> set ys \<Longrightarrow> y \<in> set ys \<Longrightarrow> P z \<Longrightarrow> x \<odot> (y \<odot> z) = y \<odot> (x \<odot> z)\<close>  | 
|
2088  | 
by (rule Cons.prems(3)) (simp_all add: *)  | 
|
2089  | 
moreover from * have \<open>x \<in> set ys\<close>  | 
|
2090  | 
by simp  | 
|
2091  | 
ultimately have \<open>fold (\<odot>) ys z = fold (\<odot>) (remove1 x ys) (x \<odot> z)\<close>  | 
|
2092  | 
by (induction ys arbitrary: z) auto  | 
|
2093  | 
moreover from Cons.prems have \<open>fold (\<odot>) xs (x \<odot> z) = fold (\<odot>) (remove1 x ys) (x \<odot> z)\<close>  | 
|
2094  | 
by (auto intro: Cons.IH)  | 
|
2095  | 
ultimately show ?case  | 
|
2096  | 
by simp  | 
|
| 
45989
 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 
haftmann 
parents: 
45866 
diff
changeset
 | 
2097  | 
qed  | 
| 
 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 
haftmann 
parents: 
45866 
diff
changeset
 | 
2098  | 
|
| 69107 | 2099  | 
lemma mset_shuffles: "zs \<in> shuffles xs ys \<Longrightarrow> mset zs = mset xs + mset ys"  | 
2100  | 
by (induction xs ys arbitrary: zs rule: shuffles.induct) auto  | 
|
| 
65350
 
b149abe619f7
added shuffle product to HOL/List
 
eberlm <eberlm@in.tum.de> 
parents: 
65048 
diff
changeset
 | 
2101  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2102  | 
lemma mset_insort [simp]: "mset (insort x xs) = add_mset x (mset xs)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2103  | 
by (induct xs) simp_all  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2104  | 
|
| 
63524
 
4ec755485732
adding mset_map to the simp rules
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63410 
diff
changeset
 | 
2105  | 
lemma mset_map[simp]: "mset (map f xs) = image_mset f (mset xs)"  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
2106  | 
by (induct xs) simp_all  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
2107  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2108  | 
global_interpretation mset_set: folding add_mset "{#}"
 | 
| 73832 | 2109  | 
  defines mset_set = "folding_on.F add_mset {#}"
 | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2110  | 
by standard (simp add: fun_eq_iff)  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2111  | 
|
| 
66276
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2112  | 
lemma sum_multiset_singleton [simp]: "sum (\<lambda>n. {#n#}) A = mset_set A"
 | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2113  | 
by (induction A rule: infinite_finite_induct) auto  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2114  | 
|
| 60513 | 2115  | 
lemma count_mset_set [simp]:  | 
2116  | 
"finite A \<Longrightarrow> x \<in> A \<Longrightarrow> count (mset_set A) x = 1" (is "PROP ?P")  | 
|
2117  | 
"\<not> finite A \<Longrightarrow> count (mset_set A) x = 0" (is "PROP ?Q")  | 
|
2118  | 
"x \<notin> A \<Longrightarrow> count (mset_set A) x = 0" (is "PROP ?R")  | 
|
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
2119  | 
proof -  | 
| 60606 | 2120  | 
have *: "count (mset_set A) x = 0" if "x \<notin> A" for A  | 
2121  | 
proof (cases "finite A")  | 
|
2122  | 
case False then show ?thesis by simp  | 
|
2123  | 
next  | 
|
2124  | 
case True from True \<open>x \<notin> A\<close> show ?thesis by (induct A) auto  | 
|
2125  | 
qed  | 
|
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
2126  | 
then show "PROP ?P" "PROP ?Q" "PROP ?R"  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
2127  | 
by (auto elim!: Set.set_insert)  | 
| 69593 | 2128  | 
qed \<comment> \<open>TODO: maybe define \<^const>\<open>mset_set\<close> also in terms of \<^const>\<open>Abs_multiset\<close>\<close>  | 
| 60513 | 2129  | 
|
2130  | 
lemma elem_mset_set[simp, intro]: "finite A \<Longrightarrow> x \<in># mset_set A \<longleftrightarrow> x \<in> A"  | 
|
| 59813 | 2131  | 
by (induct A rule: finite_induct) simp_all  | 
2132  | 
||
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2133  | 
lemma mset_set_Union:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2134  | 
  "finite A \<Longrightarrow> finite B \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> mset_set (A \<union> B) = mset_set A + mset_set B"
 | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2135  | 
by (induction A rule: finite_induct) auto  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2136  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2137  | 
lemma filter_mset_mset_set [simp]:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2138  | 
  "finite A \<Longrightarrow> filter_mset P (mset_set A) = mset_set {x\<in>A. P x}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2139  | 
proof (induction A rule: finite_induct)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2140  | 
case (insert x A)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2141  | 
from insert.hyps have "filter_mset P (mset_set (insert x A)) =  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2142  | 
      filter_mset P (mset_set A) + mset_set (if P x then {x} else {})"
 | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2143  | 
by simp  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2144  | 
  also have "filter_mset P (mset_set A) = mset_set {x\<in>A. P x}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2145  | 
by (rule insert.IH)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2146  | 
also from insert.hyps  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2147  | 
    have "\<dots> + mset_set (if P x then {x} else {}) =
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2148  | 
            mset_set ({x \<in> A. P x} \<union> (if P x then {x} else {}))" (is "_ = mset_set ?A")
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2149  | 
by (intro mset_set_Union [symmetric]) simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2150  | 
  also from insert.hyps have "?A = {y\<in>insert x A. P y}" by auto
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2151  | 
finally show ?case .  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2152  | 
qed simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2153  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2154  | 
lemma mset_set_Diff:  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2155  | 
assumes "finite A" "B \<subseteq> A"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2156  | 
shows "mset_set (A - B) = mset_set A - mset_set B"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2157  | 
proof -  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2158  | 
from assms have "mset_set ((A - B) \<union> B) = mset_set (A - B) + mset_set B"  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2159  | 
by (intro mset_set_Union) (auto dest: finite_subset)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2160  | 
also from assms have "A - B \<union> B = A" by blast  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2161  | 
finally show ?thesis by simp  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2162  | 
qed  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2163  | 
|
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2164  | 
lemma mset_set_set: "distinct xs \<Longrightarrow> mset_set (set xs) = mset xs"  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2165  | 
by (induction xs) simp_all  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2166  | 
|
| 
66276
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2167  | 
lemma count_mset_set': "count (mset_set A) x = (if finite A \<and> x \<in> A then 1 else 0)"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2168  | 
by auto  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2169  | 
|
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2170  | 
lemma subset_imp_msubset_mset_set:  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2171  | 
assumes "A \<subseteq> B" "finite B"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2172  | 
shows "mset_set A \<subseteq># mset_set B"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2173  | 
proof (rule mset_subset_eqI)  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2174  | 
fix x :: 'a  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2175  | 
from assms have "finite A" by (rule finite_subset)  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2176  | 
with assms show "count (mset_set A) x \<le> count (mset_set B) x"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2177  | 
by (cases "x \<in> A"; cases "x \<in> B") auto  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2178  | 
qed  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2179  | 
|
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2180  | 
lemma mset_set_set_mset_msubset: "mset_set (set_mset A) \<subseteq># A"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2181  | 
proof (rule mset_subset_eqI)  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2182  | 
fix x show "count (mset_set (set_mset A)) x \<le> count A x"  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2183  | 
by (cases "x \<in># A") simp_all  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2184  | 
qed  | 
| 
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2185  | 
|
| 73466 | 2186  | 
lemma mset_set_upto_eq_mset_upto:  | 
2187  | 
  \<open>mset_set {..<n} = mset [0..<n]\<close>
 | 
|
2188  | 
by (induction n) (auto simp: ac_simps lessThan_Suc)  | 
|
2189  | 
||
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2190  | 
context linorder  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2191  | 
begin  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2192  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2193  | 
definition sorted_list_of_multiset :: "'a multiset \<Rightarrow> 'a list"  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2194  | 
where  | 
| 
59998
 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 
nipkow 
parents: 
59986 
diff
changeset
 | 
2195  | 
"sorted_list_of_multiset M = fold_mset insort [] M"  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2196  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2197  | 
lemma sorted_list_of_multiset_empty [simp]:  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2198  | 
  "sorted_list_of_multiset {#} = []"
 | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2199  | 
by (simp add: sorted_list_of_multiset_def)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2200  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2201  | 
lemma sorted_list_of_multiset_singleton [simp]:  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2202  | 
  "sorted_list_of_multiset {#x#} = [x]"
 | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2203  | 
proof -  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2204  | 
interpret comp_fun_commute insort by (fact comp_fun_commute_insort)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2205  | 
show ?thesis by (simp add: sorted_list_of_multiset_def)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2206  | 
qed  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2207  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2208  | 
lemma sorted_list_of_multiset_insert [simp]:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2209  | 
"sorted_list_of_multiset (add_mset x M) = List.insort x (sorted_list_of_multiset M)"  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2210  | 
proof -  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2211  | 
interpret comp_fun_commute insort by (fact comp_fun_commute_insort)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2212  | 
show ?thesis by (simp add: sorted_list_of_multiset_def)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2213  | 
qed  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2214  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2215  | 
end  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2216  | 
|
| 66494 | 2217  | 
lemma mset_sorted_list_of_multiset[simp]: "mset (sorted_list_of_multiset M) = M"  | 
2218  | 
by (induct M) simp_all  | 
|
2219  | 
||
2220  | 
lemma sorted_list_of_multiset_mset[simp]: "sorted_list_of_multiset (mset xs) = sort xs"  | 
|
2221  | 
by (induct xs) simp_all  | 
|
2222  | 
||
2223  | 
lemma finite_set_mset_mset_set[simp]: "finite A \<Longrightarrow> set_mset (mset_set A) = A"  | 
|
2224  | 
by auto  | 
|
| 60513 | 2225  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2226  | 
lemma mset_set_empty_iff: "mset_set A = {#} \<longleftrightarrow> A = {} \<or> infinite A"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2227  | 
using finite_set_mset_mset_set by fastforce  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2228  | 
|
| 66494 | 2229  | 
lemma infinite_set_mset_mset_set: "\<not> finite A \<Longrightarrow> set_mset (mset_set A) = {}"
 | 
2230  | 
by simp  | 
|
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2231  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2232  | 
lemma set_sorted_list_of_multiset [simp]:  | 
| 60495 | 2233  | 
"set (sorted_list_of_multiset M) = set_mset M"  | 
| 
66434
 
5d7e770c7d5d
added sorted_wrt to List; added Data_Structures/Binomial_Heap.thy
 
nipkow 
parents: 
66425 
diff
changeset
 | 
2234  | 
by (induct M) (simp_all add: set_insort_key)  | 
| 60513 | 2235  | 
|
2236  | 
lemma sorted_list_of_mset_set [simp]:  | 
|
2237  | 
"sorted_list_of_multiset (mset_set A) = sorted_list_of_set A"  | 
|
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2238  | 
by (cases "finite A") (induct A rule: finite_induct, simp_all)  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2239  | 
|
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2240  | 
lemma mset_upt [simp]: "mset [m..<n] = mset_set {m..<n}"
 | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2241  | 
by (induction n) (simp_all add: atLeastLessThanSuc)  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2242  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2243  | 
lemma image_mset_map_of:  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2244  | 
  "distinct (map fst xs) \<Longrightarrow> {#the (map_of xs i). i \<in># mset (map fst xs)#} = mset (map snd xs)"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2245  | 
proof (induction xs)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2246  | 
case (Cons x xs)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2247  | 
  have "{#the (map_of (x # xs) i). i \<in># mset (map fst (x # xs))#} =
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2248  | 
          add_mset (snd x) {#the (if i = fst x then Some (snd x) else map_of xs i).
 | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2249  | 
i \<in># mset (map fst xs)#}" (is "_ = add_mset _ ?A") by simp  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2250  | 
  also from Cons.prems have "?A = {#the (map_of xs i). i :# mset (map fst xs)#}"
 | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2251  | 
by (cases x, intro image_mset_cong) (auto simp: in_multiset_in_set)  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2252  | 
also from Cons.prems have "\<dots> = mset (map snd xs)" by (intro Cons.IH) simp_all  | 
| 
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2253  | 
finally show ?case by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2254  | 
qed simp_all  | 
| 
63099
 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 
eberlm 
parents: 
63092 
diff
changeset
 | 
2255  | 
|
| 66494 | 2256  | 
lemma msubset_mset_set_iff[simp]:  | 
| 
66276
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2257  | 
assumes "finite A" "finite B"  | 
| 66494 | 2258  | 
shows "mset_set A \<subseteq># mset_set B \<longleftrightarrow> A \<subseteq> B"  | 
2259  | 
using assms set_mset_mono subset_imp_msubset_mset_set by fastforce  | 
|
2260  | 
||
2261  | 
lemma mset_set_eq_iff[simp]:  | 
|
2262  | 
assumes "finite A" "finite B"  | 
|
2263  | 
shows "mset_set A = mset_set B \<longleftrightarrow> A = B"  | 
|
2264  | 
using assms by (fastforce dest: finite_set_mset_mset_set)  | 
|
| 
66276
 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 
eberlm <eberlm@in.tum.de> 
parents: 
65547 
diff
changeset
 | 
2265  | 
|
| 
69895
 
6b03a8cf092d
more formal contributors (with the help of the history);
 
wenzelm 
parents: 
69605 
diff
changeset
 | 
2266  | 
lemma image_mset_mset_set: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close>  | 
| 
63921
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2267  | 
assumes "inj_on f A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2268  | 
shows "image_mset f (mset_set A) = mset_set (f ` A)"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2269  | 
proof cases  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2270  | 
assume "finite A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2271  | 
from this \<open>inj_on f A\<close> show ?thesis  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2272  | 
by (induct A) auto  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2273  | 
next  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2274  | 
assume "infinite A"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2275  | 
from this \<open>inj_on f A\<close> have "infinite (f ` A)"  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2276  | 
using finite_imageD by blast  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2277  | 
from \<open>infinite A\<close> \<open>infinite (f ` A)\<close> show ?thesis by simp  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2278  | 
qed  | 
| 
 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 
eberlm <eberlm@in.tum.de> 
parents: 
63908 
diff
changeset
 | 
2279  | 
|
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2280  | 
|
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
2281  | 
subsection \<open>More properties of the replicate and repeat operations\<close>  | 
| 60804 | 2282  | 
|
2283  | 
lemma in_replicate_mset[simp]: "x \<in># replicate_mset n y \<longleftrightarrow> n > 0 \<and> x = y"  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2284  | 
unfolding replicate_mset_def by (induct n) auto  | 
| 60804 | 2285  | 
|
2286  | 
lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})"
 | 
|
2287  | 
by (auto split: if_splits)  | 
|
2288  | 
||
2289  | 
lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n"  | 
|
2290  | 
by (induct n, simp_all)  | 
|
2291  | 
||
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
2292  | 
lemma count_le_replicate_mset_subset_eq: "n \<le> count M x \<longleftrightarrow> replicate_mset n x \<subseteq># M"  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
2293  | 
by (auto simp add: mset_subset_eqI) (metis count_replicate_mset subseteq_mset_def)  | 
| 60804 | 2294  | 
|
2295  | 
lemma filter_eq_replicate_mset: "{#y \<in># D. y = x#} = replicate_mset (count D x) x"
 | 
|
2296  | 
by (induct D) simp_all  | 
|
2297  | 
||
| 66494 | 2298  | 
lemma replicate_count_mset_eq_filter_eq: "replicate (count (mset xs) k) k = filter (HOL.eq k) xs"  | 
| 
61031
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2299  | 
by (induct xs) auto  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2300  | 
|
| 66494 | 2301  | 
lemma replicate_mset_eq_empty_iff [simp]: "replicate_mset n a = {#} \<longleftrightarrow> n = 0"
 | 
| 62366 | 2302  | 
by (induct n) simp_all  | 
2303  | 
||
2304  | 
lemma replicate_mset_eq_iff:  | 
|
| 66494 | 2305  | 
"replicate_mset m a = replicate_mset n b \<longleftrightarrow> m = 0 \<and> n = 0 \<or> m = n \<and> a = b"  | 
| 62366 | 2306  | 
by (auto simp add: multiset_eq_iff)  | 
2307  | 
||
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
2308  | 
lemma repeat_mset_cancel1: "repeat_mset a A = repeat_mset a B \<longleftrightarrow> A = B \<or> a = 0"  | 
| 
63849
 
0dd6731060d7
delete looping simp rule
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63831 
diff
changeset
 | 
2309  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
0dd6731060d7
delete looping simp rule
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63831 
diff
changeset
 | 
2310  | 
|
| 
63908
 
ca41b6670904
support replicate_mset in multiset simproc
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63882 
diff
changeset
 | 
2311  | 
lemma repeat_mset_cancel2: "repeat_mset a A = repeat_mset b A \<longleftrightarrow> a = b \<or> A = {#}"
 | 
| 
63849
 
0dd6731060d7
delete looping simp rule
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63831 
diff
changeset
 | 
2312  | 
by (auto simp: multiset_eq_iff)  | 
| 
 
0dd6731060d7
delete looping simp rule
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63831 
diff
changeset
 | 
2313  | 
|
| 64077 | 2314  | 
lemma repeat_mset_eq_empty_iff: "repeat_mset n A = {#} \<longleftrightarrow> n = 0 \<or> A = {#}"
 | 
2315  | 
by (cases n) auto  | 
|
2316  | 
||
| 63924 | 2317  | 
lemma image_replicate_mset [simp]:  | 
2318  | 
"image_mset f (replicate_mset n a) = replicate_mset n (f a)"  | 
|
2319  | 
by (induct n) simp_all  | 
|
2320  | 
||
| 67051 | 2321  | 
lemma replicate_mset_msubseteq_iff:  | 
2322  | 
"replicate_mset m a \<subseteq># replicate_mset n b \<longleftrightarrow> m = 0 \<or> a = b \<and> m \<le> n"  | 
|
2323  | 
by (cases m)  | 
|
| 68406 | 2324  | 
(auto simp: insert_subset_eq_iff simp flip: count_le_replicate_mset_subset_eq)  | 
| 67051 | 2325  | 
|
2326  | 
lemma msubseteq_replicate_msetE:  | 
|
2327  | 
assumes "A \<subseteq># replicate_mset n a"  | 
|
2328  | 
obtains m where "m \<le> n" and "A = replicate_mset m a"  | 
|
2329  | 
proof (cases "n = 0")  | 
|
2330  | 
case True  | 
|
2331  | 
with assms that show thesis  | 
|
2332  | 
by simp  | 
|
2333  | 
next  | 
|
2334  | 
case False  | 
|
2335  | 
from assms have "set_mset A \<subseteq> set_mset (replicate_mset n a)"  | 
|
2336  | 
by (rule set_mset_mono)  | 
|
2337  | 
  with False have "set_mset A \<subseteq> {a}"
 | 
|
2338  | 
by simp  | 
|
2339  | 
then have "\<exists>m. A = replicate_mset m a"  | 
|
2340  | 
proof (induction A)  | 
|
2341  | 
case empty  | 
|
2342  | 
then show ?case  | 
|
2343  | 
by simp  | 
|
2344  | 
next  | 
|
2345  | 
case (add b A)  | 
|
2346  | 
then obtain m where "A = replicate_mset m a"  | 
|
2347  | 
by auto  | 
|
2348  | 
with add.prems show ?case  | 
|
2349  | 
by (auto intro: exI [of _ "Suc m"])  | 
|
2350  | 
qed  | 
|
2351  | 
then obtain m where A: "A = replicate_mset m a" ..  | 
|
2352  | 
with assms have "m \<le> n"  | 
|
2353  | 
by (auto simp add: replicate_mset_msubseteq_iff)  | 
|
2354  | 
then show thesis using A ..  | 
|
2355  | 
qed  | 
|
2356  | 
||
| 60804 | 2357  | 
|
| 60500 | 2358  | 
subsection \<open>Big operators\<close>  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2359  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2360  | 
locale comm_monoid_mset = comm_monoid  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2361  | 
begin  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2362  | 
|
| 64075 | 2363  | 
interpretation comp_fun_commute f  | 
2364  | 
by standard (simp add: fun_eq_iff left_commute)  | 
|
2365  | 
||
2366  | 
interpretation comp?: comp_fun_commute "f \<circ> g"  | 
|
2367  | 
by (fact comp_comp_fun_commute)  | 
|
2368  | 
||
2369  | 
context  | 
|
2370  | 
begin  | 
|
2371  | 
||
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2372  | 
definition F :: "'a multiset \<Rightarrow> 'a"  | 
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63195 
diff
changeset
 | 
2373  | 
where eq_fold: "F M = fold_mset f \<^bold>1 M"  | 
| 
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63195 
diff
changeset
 | 
2374  | 
|
| 
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63195 
diff
changeset
 | 
2375  | 
lemma empty [simp]: "F {#} = \<^bold>1"
 | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2376  | 
by (simp add: eq_fold)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2377  | 
|
| 60678 | 2378  | 
lemma singleton [simp]: "F {#x#} = x"
 | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2379  | 
proof -  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2380  | 
interpret comp_fun_commute  | 
| 60678 | 2381  | 
by standard (simp add: fun_eq_iff left_commute)  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2382  | 
show ?thesis by (simp add: eq_fold)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2383  | 
qed  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2384  | 
|
| 
63290
 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 
haftmann 
parents: 
63195 
diff
changeset
 | 
2385  | 
lemma union [simp]: "F (M + N) = F M \<^bold>* F N"  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2386  | 
proof -  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2387  | 
interpret comp_fun_commute f  | 
| 60678 | 2388  | 
by standard (simp add: fun_eq_iff left_commute)  | 
2389  | 
show ?thesis  | 
|
2390  | 
by (induct N) (simp_all add: left_commute eq_fold)  | 
|
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2391  | 
qed  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2392  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2393  | 
lemma add_mset [simp]: "F (add_mset x N) = x \<^bold>* F N"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2394  | 
unfolding add_mset_add_single[of x N] union by (simp add: ac_simps)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2395  | 
|
| 64075 | 2396  | 
lemma insert [simp]:  | 
2397  | 
shows "F (image_mset g (add_mset x A)) = g x \<^bold>* F (image_mset g A)"  | 
|
2398  | 
by (simp add: eq_fold)  | 
|
2399  | 
||
2400  | 
lemma remove:  | 
|
2401  | 
assumes "x \<in># A"  | 
|
2402  | 
  shows "F A = x \<^bold>* F (A - {#x#})"
 | 
|
2403  | 
using multi_member_split[OF assms] by auto  | 
|
2404  | 
||
2405  | 
lemma neutral:  | 
|
2406  | 
"\<forall>x\<in>#A. x = \<^bold>1 \<Longrightarrow> F A = \<^bold>1"  | 
|
2407  | 
by (induct A) simp_all  | 
|
2408  | 
||
2409  | 
lemma neutral_const [simp]:  | 
|
2410  | 
"F (image_mset (\<lambda>_. \<^bold>1) A) = \<^bold>1"  | 
|
2411  | 
by (simp add: neutral)  | 
|
2412  | 
||
2413  | 
private lemma F_image_mset_product:  | 
|
2414  | 
  "F {#g x j \<^bold>* F {#g i j. i \<in># A#}. j \<in># B#} =
 | 
|
2415  | 
    F (image_mset (g x) B) \<^bold>* F {#F {#g i j. i \<in># A#}. j \<in># B#}"
 | 
|
2416  | 
by (induction B) (simp_all add: left_commute semigroup.assoc semigroup_axioms)  | 
|
2417  | 
||
| 68938 | 2418  | 
lemma swap:  | 
| 64075 | 2419  | 
"F (image_mset (\<lambda>i. F (image_mset (g i) B)) A) =  | 
2420  | 
F (image_mset (\<lambda>j. F (image_mset (\<lambda>i. g i j) A)) B)"  | 
|
2421  | 
apply (induction A, simp)  | 
|
2422  | 
apply (induction B, auto simp add: F_image_mset_product ac_simps)  | 
|
2423  | 
done  | 
|
2424  | 
||
2425  | 
lemma distrib: "F (image_mset (\<lambda>x. g x \<^bold>* h x) A) = F (image_mset g A) \<^bold>* F (image_mset h A)"  | 
|
2426  | 
by (induction A) (auto simp: ac_simps)  | 
|
2427  | 
||
2428  | 
lemma union_disjoint:  | 
|
2429  | 
  "A \<inter># B = {#} \<Longrightarrow> F (image_mset g (A \<union># B)) = F (image_mset g A) \<^bold>* F (image_mset g B)"
 | 
|
2430  | 
by (induction A) (auto simp: ac_simps)  | 
|
2431  | 
||
2432  | 
end  | 
|
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2433  | 
end  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2434  | 
|
| 67398 | 2435  | 
lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute ((+) :: 'a multiset \<Rightarrow> _ \<Rightarrow> _)"  | 
| 60678 | 2436  | 
by standard (simp add: add_ac comp_def)  | 
| 59813 | 2437  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2438  | 
declare comp_fun_commute.fold_mset_add_mset[OF comp_fun_commute_plus_mset, simp]  | 
| 59813 | 2439  | 
|
| 67398 | 2440  | 
lemma in_mset_fold_plus_iff[iff]: "x \<in># fold_mset (+) M NN \<longleftrightarrow> x \<in># M \<or> (\<exists>N. N \<in># NN \<and> x \<in># N)"  | 
| 59813 | 2441  | 
by (induct NN) auto  | 
2442  | 
||
| 54868 | 2443  | 
context comm_monoid_add  | 
2444  | 
begin  | 
|
2445  | 
||
| 63830 | 2446  | 
sublocale sum_mset: comm_monoid_mset plus 0  | 
2447  | 
defines sum_mset = sum_mset.F ..  | 
|
2448  | 
||
| 64267 | 2449  | 
lemma sum_unfold_sum_mset:  | 
2450  | 
"sum f A = sum_mset (image_mset f (mset_set A))"  | 
|
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2451  | 
by (cases "finite A") (induct A rule: finite_induct, simp_all)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2452  | 
|
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2453  | 
end  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2454  | 
|
| 
73047
 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72607 
diff
changeset
 | 
2455  | 
notation sum_mset ("\<Sum>\<^sub>#")
 | 
| 
 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72607 
diff
changeset
 | 
2456  | 
|
| 62366 | 2457  | 
syntax (ASCII)  | 
| 63830 | 2458  | 
  "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3SUM _:#_. _)" [0, 51, 10] 10)
 | 
| 62366 | 2459  | 
syntax  | 
| 63830 | 2460  | 
  "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 62366 | 2461  | 
translations  | 
| 63830 | 2462  | 
"\<Sum>i \<in># A. b" \<rightleftharpoons> "CONST sum_mset (CONST image_mset (\<lambda>i. b) A)"  | 
| 59949 | 2463  | 
|
| 66938 | 2464  | 
context comm_monoid_add  | 
2465  | 
begin  | 
|
2466  | 
||
2467  | 
lemma sum_mset_sum_list:  | 
|
2468  | 
"sum_mset (mset xs) = sum_list xs"  | 
|
2469  | 
by (induction xs) auto  | 
|
2470  | 
||
2471  | 
end  | 
|
2472  | 
||
2473  | 
context canonically_ordered_monoid_add  | 
|
2474  | 
begin  | 
|
2475  | 
||
2476  | 
lemma sum_mset_0_iff [simp]:  | 
|
2477  | 
"sum_mset M = 0 \<longleftrightarrow> (\<forall>x \<in> set_mset M. x = 0)"  | 
|
2478  | 
by (induction M) auto  | 
|
2479  | 
||
2480  | 
end  | 
|
2481  | 
||
2482  | 
context ordered_comm_monoid_add  | 
|
2483  | 
begin  | 
|
2484  | 
||
2485  | 
lemma sum_mset_mono:  | 
|
2486  | 
"sum_mset (image_mset f K) \<le> sum_mset (image_mset g K)"  | 
|
2487  | 
if "\<And>i. i \<in># K \<Longrightarrow> f i \<le> g i"  | 
|
2488  | 
using that by (induction K) (simp_all add: add_mono)  | 
|
2489  | 
||
2490  | 
end  | 
|
2491  | 
||
| 73470 | 2492  | 
context cancel_comm_monoid_add  | 
| 66938 | 2493  | 
begin  | 
2494  | 
||
2495  | 
lemma sum_mset_diff:  | 
|
2496  | 
"sum_mset (M - N) = sum_mset M - sum_mset N" if "N \<subseteq># M" for M N :: "'a multiset"  | 
|
2497  | 
using that by (auto simp add: subset_mset.le_iff_add)  | 
|
2498  | 
||
2499  | 
end  | 
|
2500  | 
||
2501  | 
context semiring_0  | 
|
2502  | 
begin  | 
|
2503  | 
||
| 63860 | 2504  | 
lemma sum_mset_distrib_left:  | 
| 66938 | 2505  | 
"c * (\<Sum>x \<in># M. f x) = (\<Sum>x \<in># M. c * f(x))"  | 
2506  | 
by (induction M) (simp_all add: algebra_simps)  | 
|
| 63860 | 2507  | 
|
| 64075 | 2508  | 
lemma sum_mset_distrib_right:  | 
| 66938 | 2509  | 
"(\<Sum>x \<in># M. f x) * c = (\<Sum>x \<in># M. f x * c)"  | 
2510  | 
by (induction M) (simp_all add: algebra_simps)  | 
|
2511  | 
||
2512  | 
end  | 
|
2513  | 
||
2514  | 
lemma sum_mset_product:  | 
|
2515  | 
  fixes f :: "'a::{comm_monoid_add,times} \<Rightarrow> 'b::semiring_0"
 | 
|
2516  | 
shows "(\<Sum>i \<in># A. f i) * (\<Sum>i \<in># B. g i) = (\<Sum>i\<in>#A. \<Sum>j\<in>#B. f i * g j)"  | 
|
| 68938 | 2517  | 
by (subst sum_mset.swap) (simp add: sum_mset_distrib_left sum_mset_distrib_right)  | 
| 66938 | 2518  | 
|
2519  | 
context semiring_1  | 
|
2520  | 
begin  | 
|
2521  | 
||
2522  | 
lemma sum_mset_replicate_mset [simp]:  | 
|
2523  | 
"sum_mset (replicate_mset n a) = of_nat n * a"  | 
|
2524  | 
by (induction n) (simp_all add: algebra_simps)  | 
|
2525  | 
||
2526  | 
lemma sum_mset_delta:  | 
|
2527  | 
"sum_mset (image_mset (\<lambda>x. if x = y then c else 0) A) = c * of_nat (count A y)"  | 
|
2528  | 
by (induction A) (simp_all add: algebra_simps)  | 
|
2529  | 
||
2530  | 
lemma sum_mset_delta':  | 
|
2531  | 
"sum_mset (image_mset (\<lambda>x. if y = x then c else 0) A) = c * of_nat (count A y)"  | 
|
2532  | 
by (induction A) (simp_all add: algebra_simps)  | 
|
2533  | 
||
2534  | 
end  | 
|
2535  | 
||
2536  | 
lemma of_nat_sum_mset [simp]:  | 
|
2537  | 
"of_nat (sum_mset A) = sum_mset (image_mset of_nat A)"  | 
|
2538  | 
by (induction A) auto  | 
|
2539  | 
||
2540  | 
lemma size_eq_sum_mset:  | 
|
2541  | 
"size M = (\<Sum>a\<in>#M. 1)"  | 
|
2542  | 
using image_mset_const_eq [of "1::nat" M] by simp  | 
|
2543  | 
||
2544  | 
lemma size_mset_set [simp]:  | 
|
2545  | 
"size (mset_set A) = card A"  | 
|
2546  | 
by (simp only: size_eq_sum_mset card_eq_sum sum_unfold_sum_mset)  | 
|
| 64075 | 2547  | 
|
2548  | 
lemma sum_mset_constant [simp]:  | 
|
2549  | 
fixes y :: "'b::semiring_1"  | 
|
2550  | 
shows \<open>(\<Sum>x\<in>#A. y) = of_nat (size A) * y\<close>  | 
|
2551  | 
by (induction A) (auto simp: algebra_simps)  | 
|
2552  | 
||
| 
73047
 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72607 
diff
changeset
 | 
2553  | 
lemma set_mset_Union_mset[simp]: "set_mset (\<Sum>\<^sub># MM) = (\<Union>M \<in> set_mset MM. set_mset M)"  | 
| 59813 | 2554  | 
by (induct MM) auto  | 
2555  | 
||
| 
73047
 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72607 
diff
changeset
 | 
2556  | 
lemma in_Union_mset_iff[iff]: "x \<in># \<Sum>\<^sub># MM \<longleftrightarrow> (\<exists>M. M \<in># MM \<and> x \<in># M)"  | 
| 59813 | 2557  | 
by (induct MM) auto  | 
2558  | 
||
| 64267 | 2559  | 
lemma count_sum:  | 
2560  | 
"count (sum f A) x = sum (\<lambda>a. count (f a) x) A"  | 
|
| 62366 | 2561  | 
by (induct A rule: infinite_finite_induct) simp_all  | 
2562  | 
||
| 64267 | 2563  | 
lemma sum_eq_empty_iff:  | 
| 62366 | 2564  | 
assumes "finite A"  | 
| 64267 | 2565  | 
  shows "sum f A = {#} \<longleftrightarrow> (\<forall>a\<in>A. f a = {#})"
 | 
| 62366 | 2566  | 
using assms by induct simp_all  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2567  | 
|
| 
73047
 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72607 
diff
changeset
 | 
2568  | 
lemma Union_mset_empty_conv[simp]: "\<Sum>\<^sub># M = {#} \<longleftrightarrow> (\<forall>i\<in>#M. i = {#})"
 | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
2569  | 
by (induction M) auto  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
2570  | 
|
| 
73047
 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
72607 
diff
changeset
 | 
2571  | 
lemma Union_image_single_mset[simp]: "\<Sum>\<^sub># (image_mset (\<lambda>x. {#x#}) m) = m"
 | 
| 67656 | 2572  | 
by(induction m) auto  | 
2573  | 
||
| 66938 | 2574  | 
|
| 54868 | 2575  | 
context comm_monoid_mult  | 
2576  | 
begin  | 
|
2577  | 
||
| 63830 | 2578  | 
sublocale prod_mset: comm_monoid_mset times 1  | 
2579  | 
defines prod_mset = prod_mset.F ..  | 
|
2580  | 
||
2581  | 
lemma prod_mset_empty:  | 
|
2582  | 
  "prod_mset {#} = 1"
 | 
|
2583  | 
by (fact prod_mset.empty)  | 
|
2584  | 
||
2585  | 
lemma prod_mset_singleton:  | 
|
2586  | 
  "prod_mset {#x#} = x"
 | 
|
2587  | 
by (fact prod_mset.singleton)  | 
|
2588  | 
||
2589  | 
lemma prod_mset_Un:  | 
|
2590  | 
"prod_mset (A + B) = prod_mset A * prod_mset B"  | 
|
2591  | 
by (fact prod_mset.union)  | 
|
2592  | 
||
| 66938 | 2593  | 
lemma prod_mset_prod_list:  | 
2594  | 
"prod_mset (mset xs) = prod_list xs"  | 
|
2595  | 
by (induct xs) auto  | 
|
2596  | 
||
| 63830 | 2597  | 
lemma prod_mset_replicate_mset [simp]:  | 
2598  | 
"prod_mset (replicate_mset n a) = a ^ n"  | 
|
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2599  | 
by (induct n) simp_all  | 
| 60804 | 2600  | 
|
| 64272 | 2601  | 
lemma prod_unfold_prod_mset:  | 
2602  | 
"prod f A = prod_mset (image_mset f (mset_set A))"  | 
|
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2603  | 
by (cases "finite A") (induct A rule: finite_induct, simp_all)  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2604  | 
|
| 63830 | 2605  | 
lemma prod_mset_multiplicity:  | 
| 64272 | 2606  | 
"prod_mset M = prod (\<lambda>x. x ^ count M x) (set_mset M)"  | 
2607  | 
by (simp add: fold_mset_def prod.eq_fold prod_mset.eq_fold funpow_times_power comp_def)  | 
|
| 63830 | 2608  | 
|
2609  | 
lemma prod_mset_delta: "prod_mset (image_mset (\<lambda>x. if x = y then c else 1) A) = c ^ count A y"  | 
|
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2610  | 
by (induction A) simp_all  | 
| 
63534
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
63524 
diff
changeset
 | 
2611  | 
|
| 63830 | 2612  | 
lemma prod_mset_delta': "prod_mset (image_mset (\<lambda>x. if y = x then c else 1) A) = c ^ count A y"  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2613  | 
by (induction A) simp_all  | 
| 
63534
 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 
eberlm <eberlm@in.tum.de> 
parents: 
63524 
diff
changeset
 | 
2614  | 
|
| 66938 | 2615  | 
lemma prod_mset_subset_imp_dvd:  | 
2616  | 
assumes "A \<subseteq># B"  | 
|
2617  | 
shows "prod_mset A dvd prod_mset B"  | 
|
2618  | 
proof -  | 
|
2619  | 
from assms have "B = (B - A) + A" by (simp add: subset_mset.diff_add)  | 
|
2620  | 
also have "prod_mset \<dots> = prod_mset (B - A) * prod_mset A" by simp  | 
|
2621  | 
also have "prod_mset A dvd \<dots>" by simp  | 
|
2622  | 
finally show ?thesis .  | 
|
2623  | 
qed  | 
|
2624  | 
||
2625  | 
lemma dvd_prod_mset:  | 
|
2626  | 
assumes "x \<in># A"  | 
|
2627  | 
shows "x dvd prod_mset A"  | 
|
2628  | 
  using assms prod_mset_subset_imp_dvd [of "{#x#}" A] by simp
 | 
|
2629  | 
||
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2630  | 
end  | 
| 
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2631  | 
|
| 
73052
 
c03a148110cc
HOL-Library.Multiset: new notation for prod_mset, consistent with sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
73047 
diff
changeset
 | 
2632  | 
notation prod_mset ("\<Prod>\<^sub>#")
 | 
| 
 
c03a148110cc
HOL-Library.Multiset: new notation for prod_mset, consistent with sum_mset
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
73047 
diff
changeset
 | 
2633  | 
|
| 
61955
 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 
wenzelm 
parents: 
61890 
diff
changeset
 | 
2634  | 
syntax (ASCII)  | 
| 63830 | 2635  | 
  "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3PROD _:#_. _)" [0, 51, 10] 10)
 | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2636  | 
syntax  | 
| 63830 | 2637  | 
  "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3\<Prod>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2638  | 
translations  | 
| 63830 | 2639  | 
"\<Prod>i \<in># A. b" \<rightleftharpoons> "CONST prod_mset (CONST image_mset (\<lambda>i. b) A)"  | 
2640  | 
||
| 
64591
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64587 
diff
changeset
 | 
2641  | 
lemma prod_mset_constant [simp]: "(\<Prod>_\<in>#A. c) = c ^ size A"  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64587 
diff
changeset
 | 
2642  | 
by (simp add: image_mset_const_eq)  | 
| 
 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 
haftmann 
parents: 
64587 
diff
changeset
 | 
2643  | 
|
| 63830 | 2644  | 
lemma (in semidom) prod_mset_zero_iff [iff]:  | 
2645  | 
"prod_mset A = 0 \<longleftrightarrow> 0 \<in># A"  | 
|
| 62366 | 2646  | 
by (induct A) auto  | 
2647  | 
||
| 63830 | 2648  | 
lemma (in semidom_divide) prod_mset_diff:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2649  | 
assumes "B \<subseteq># A" and "0 \<notin># B"  | 
| 63830 | 2650  | 
shows "prod_mset (A - B) = prod_mset A div prod_mset B"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2651  | 
proof -  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2652  | 
from assms obtain C where "A = B + C"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2653  | 
by (metis subset_mset.add_diff_inverse)  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2654  | 
with assms show ?thesis by simp  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2655  | 
qed  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2656  | 
|
| 63830 | 2657  | 
lemma (in semidom_divide) prod_mset_minus:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2658  | 
assumes "a \<in># A" and "a \<noteq> 0"  | 
| 63830 | 2659  | 
  shows "prod_mset (A - {#a#}) = prod_mset A div a"
 | 
2660  | 
  using assms prod_mset_diff [of "{#a#}" A] by auto
 | 
|
2661  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2662  | 
lemma (in normalization_semidom) normalize_prod_mset_normalize:  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2663  | 
"normalize (prod_mset (image_mset normalize A)) = normalize (prod_mset A)"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2664  | 
proof (induction A)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2665  | 
case (add x A)  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2666  | 
have "normalize (prod_mset (image_mset normalize (add_mset x A))) =  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2667  | 
normalize (x * normalize (prod_mset (image_mset normalize A)))"  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2668  | 
by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2669  | 
also note add.IH  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2670  | 
finally show ?case by simp  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2671  | 
qed auto  | 
| 
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2672  | 
|
| 63924 | 2673  | 
lemma (in algebraic_semidom) is_unit_prod_mset_iff:  | 
2674  | 
"is_unit (prod_mset A) \<longleftrightarrow> (\<forall>x \<in># A. is_unit x)"  | 
|
2675  | 
by (induct A) (auto simp: is_unit_mult_iff)  | 
|
2676  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2677  | 
lemma (in normalization_semidom_multiplicative) normalize_prod_mset:  | 
| 63924 | 2678  | 
"normalize (prod_mset A) = prod_mset (image_mset normalize A)"  | 
2679  | 
by (induct A) (simp_all add: normalize_mult)  | 
|
2680  | 
||
| 
71398
 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
69895 
diff
changeset
 | 
2681  | 
lemma (in normalization_semidom_multiplicative) normalized_prod_msetI:  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2682  | 
assumes "\<And>a. a \<in># A \<Longrightarrow> normalize a = a"  | 
| 63830 | 2683  | 
shows "normalize (prod_mset A) = prod_mset A"  | 
| 63924 | 2684  | 
proof -  | 
2685  | 
from assms have "image_mset normalize A = A"  | 
|
2686  | 
by (induct A) simp_all  | 
|
2687  | 
then show ?thesis by (simp add: normalize_prod_mset)  | 
|
2688  | 
qed  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2689  | 
|
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2690  | 
|
| 73301 | 2691  | 
subsection \<open>Multiset as order-ignorant lists\<close>  | 
| 
51548
 
757fa47af981
centralized various multiset operations in theory multiset;
 
haftmann 
parents: 
51161 
diff
changeset
 | 
2692  | 
|
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2693  | 
context linorder  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2694  | 
begin  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2695  | 
|
| 60515 | 2696  | 
lemma mset_insort [simp]:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2697  | 
"mset (insort_key k x xs) = add_mset x (mset xs)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2698  | 
by (induct xs) simp_all  | 
| 58425 | 2699  | 
|
| 60515 | 2700  | 
lemma mset_sort [simp]:  | 
2701  | 
"mset (sort_key k xs) = mset xs"  | 
|
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2702  | 
by (induct xs) simp_all  | 
| 37107 | 2703  | 
|
| 60500 | 2704  | 
text \<open>  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2705  | 
This lemma shows which properties suffice to show that a function  | 
| 61585 | 2706  | 
\<open>f\<close> with \<open>f xs = ys\<close> behaves like sort.  | 
| 60500 | 2707  | 
\<close>  | 
| 37074 | 2708  | 
|
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2709  | 
lemma properties_for_sort_key:  | 
| 60515 | 2710  | 
assumes "mset ys = mset xs"  | 
| 60606 | 2711  | 
and "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>x. f k = f x) ys = filter (\<lambda>x. f k = f x) xs"  | 
2712  | 
and "sorted (map f ys)"  | 
|
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2713  | 
shows "sort_key f xs = ys"  | 
| 60606 | 2714  | 
using assms  | 
| 46921 | 2715  | 
proof (induct xs arbitrary: ys)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2716  | 
case Nil then show ?case by simp  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2717  | 
next  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2718  | 
case (Cons x xs)  | 
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2719  | 
from Cons.prems(2) have  | 
| 
40305
 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 
haftmann 
parents: 
40303 
diff
changeset
 | 
2720  | 
"\<forall>k \<in> set ys. filter (\<lambda>x. f k = f x) (remove1 x ys) = filter (\<lambda>x. f k = f x) xs"  | 
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2721  | 
by (simp add: filter_remove1)  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2722  | 
with Cons.prems have "sort_key f xs = remove1 x ys"  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2723  | 
by (auto intro!: Cons.hyps simp add: sorted_map_remove1)  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2724  | 
moreover from Cons.prems have "x \<in># mset ys"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2725  | 
by auto  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2726  | 
then have "x \<in> set ys"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2727  | 
by simp  | 
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2728  | 
ultimately show ?case using Cons.prems by (simp add: insort_key_remove1)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2729  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2730  | 
|
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2731  | 
lemma properties_for_sort:  | 
| 60515 | 2732  | 
assumes multiset: "mset ys = mset xs"  | 
| 60606 | 2733  | 
and "sorted ys"  | 
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2734  | 
shows "sort xs = ys"  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2735  | 
proof (rule properties_for_sort_key)  | 
| 60515 | 2736  | 
from multiset show "mset ys = mset xs" .  | 
| 60500 | 2737  | 
from \<open>sorted ys\<close> show "sorted (map (\<lambda>x. x) ys)" by simp  | 
| 60678 | 2738  | 
from multiset have "length (filter (\<lambda>y. k = y) ys) = length (filter (\<lambda>x. k = x) xs)" for k  | 
| 60515 | 2739  | 
by (rule mset_eq_length_filter)  | 
| 60678 | 2740  | 
then have "replicate (length (filter (\<lambda>y. k = y) ys)) k =  | 
2741  | 
replicate (length (filter (\<lambda>x. k = x) xs)) k" for k  | 
|
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2742  | 
by simp  | 
| 60678 | 2743  | 
then show "k \<in> set ys \<Longrightarrow> filter (\<lambda>y. k = y) ys = filter (\<lambda>x. k = x) xs" for k  | 
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2744  | 
by (simp add: replicate_length_filter)  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2745  | 
qed  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2746  | 
|
| 
61031
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2747  | 
lemma sort_key_inj_key_eq:  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2748  | 
assumes mset_equal: "mset xs = mset ys"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2749  | 
and "inj_on f (set xs)"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2750  | 
and "sorted (map f ys)"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2751  | 
shows "sort_key f xs = ys"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2752  | 
proof (rule properties_for_sort_key)  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2753  | 
from mset_equal  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2754  | 
show "mset ys = mset xs" by simp  | 
| 61188 | 2755  | 
from \<open>sorted (map f ys)\<close>  | 
| 
61031
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2756  | 
show "sorted (map f ys)" .  | 
| 68386 | 2757  | 
show "[x\<leftarrow>ys . f k = f x] = [x\<leftarrow>xs . f k = f x]" if "k \<in> set ys" for k  | 
| 
61031
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2758  | 
proof -  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2759  | 
from mset_equal  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2760  | 
have set_equal: "set xs = set ys" by (rule mset_eq_setD)  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2761  | 
with that have "insert k (set ys) = set ys" by auto  | 
| 61188 | 2762  | 
with \<open>inj_on f (set xs)\<close> have inj: "inj_on f (insert k (set ys))"  | 
| 
61031
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2763  | 
by (simp add: set_equal)  | 
| 68386 | 2764  | 
from inj have "[x\<leftarrow>ys . f k = f x] = filter (HOL.eq k) ys"  | 
| 
61031
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2765  | 
by (auto intro!: inj_on_filter_key_eq)  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2766  | 
also have "\<dots> = replicate (count (mset ys) k) k"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2767  | 
by (simp add: replicate_count_mset_eq_filter_eq)  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2768  | 
also have "\<dots> = replicate (count (mset xs) k) k"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2769  | 
using mset_equal by simp  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2770  | 
also have "\<dots> = filter (HOL.eq k) xs"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2771  | 
by (simp add: replicate_count_mset_eq_filter_eq)  | 
| 68386 | 2772  | 
also have "\<dots> = [x\<leftarrow>xs . f k = f x]"  | 
| 
61031
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2773  | 
using inj by (auto intro!: inj_on_filter_key_eq [symmetric] simp add: set_equal)  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2774  | 
finally show ?thesis .  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2775  | 
qed  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2776  | 
qed  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2777  | 
|
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2778  | 
lemma sort_key_eq_sort_key:  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2779  | 
assumes "mset xs = mset ys"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2780  | 
and "inj_on f (set xs)"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2781  | 
shows "sort_key f xs = sort_key f ys"  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2782  | 
by (rule sort_key_inj_key_eq) (simp_all add: assms)  | 
| 
 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 
haftmann 
parents: 
60804 
diff
changeset
 | 
2783  | 
|
| 
40303
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2784  | 
lemma sort_key_by_quicksort:  | 
| 68386 | 2785  | 
"sort_key f xs = sort_key f [x\<leftarrow>xs. f x < f (xs ! (length xs div 2))]  | 
2786  | 
@ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))]  | 
|
2787  | 
@ sort_key f [x\<leftarrow>xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs")  | 
|
| 
40303
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2788  | 
proof (rule properties_for_sort_key)  | 
| 60515 | 2789  | 
show "mset ?rhs = mset ?lhs"  | 
| 69442 | 2790  | 
by (rule multiset_eqI) auto  | 
| 
40303
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2791  | 
show "sorted (map f ?rhs)"  | 
| 
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2792  | 
by (auto simp add: sorted_append intro: sorted_map_same)  | 
| 
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2793  | 
next  | 
| 
40305
 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 
haftmann 
parents: 
40303 
diff
changeset
 | 
2794  | 
fix l  | 
| 
 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 
haftmann 
parents: 
40303 
diff
changeset
 | 
2795  | 
assume "l \<in> set ?rhs"  | 
| 40346 | 2796  | 
let ?pivot = "f (xs ! (length xs div 2))"  | 
2797  | 
have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto  | 
|
| 68386 | 2798  | 
have "[x \<leftarrow> sort_key f xs . f x = f l] = [x \<leftarrow> xs. f x = f l]"  | 
| 
40305
 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 
haftmann 
parents: 
40303 
diff
changeset
 | 
2799  | 
unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same)  | 
| 68386 | 2800  | 
with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp  | 
| 40346 | 2801  | 
have "\<And>x P. P (f x) ?pivot \<and> f l = f x \<longleftrightarrow> P (f l) ?pivot \<and> f l = f x" by auto  | 
| 68386 | 2802  | 
then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] =  | 
2803  | 
[x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp  | 
|
| 67398 | 2804  | 
note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"]  | 
| 68386 | 2805  | 
show "[x \<leftarrow> ?rhs. f l = f x] = [x \<leftarrow> ?lhs. f l = f x]"  | 
| 
40305
 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 
haftmann 
parents: 
40303 
diff
changeset
 | 
2806  | 
proof (cases "f l" ?pivot rule: linorder_cases)  | 
| 46730 | 2807  | 
case less  | 
2808  | 
then have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto  | 
|
2809  | 
with less show ?thesis  | 
|
| 40346 | 2810  | 
by (simp add: filter_sort [symmetric] ** ***)  | 
| 
40305
 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 
haftmann 
parents: 
40303 
diff
changeset
 | 
2811  | 
next  | 
| 40306 | 2812  | 
case equal then show ?thesis  | 
| 40346 | 2813  | 
by (simp add: * less_le)  | 
| 
40305
 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 
haftmann 
parents: 
40303 
diff
changeset
 | 
2814  | 
next  | 
| 46730 | 2815  | 
case greater  | 
2816  | 
then have "f l \<noteq> ?pivot" and "\<not> f l < ?pivot" by auto  | 
|
2817  | 
with greater show ?thesis  | 
|
| 40346 | 2818  | 
by (simp add: filter_sort [symmetric] ** ***)  | 
| 40306 | 2819  | 
qed  | 
| 
40303
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2820  | 
qed  | 
| 
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2821  | 
|
| 
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2822  | 
lemma sort_by_quicksort:  | 
| 68386 | 2823  | 
"sort xs = sort [x\<leftarrow>xs. x < xs ! (length xs div 2)]  | 
2824  | 
@ [x\<leftarrow>xs. x = xs ! (length xs div 2)]  | 
|
2825  | 
@ sort [x\<leftarrow>xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs")  | 
|
| 
40303
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2826  | 
using sort_key_by_quicksort [of "\<lambda>x. x", symmetric] by simp  | 
| 
 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 
haftmann 
parents: 
40250 
diff
changeset
 | 
2827  | 
|
| 68983 | 2828  | 
text \<open>A stable parameterized quicksort\<close>  | 
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2829  | 
|
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2830  | 
definition part :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'b list \<times> 'b list \<times> 'b list" where
 | 
| 68386 | 2831  | 
"part f pivot xs = ([x \<leftarrow> xs. f x < pivot], [x \<leftarrow> xs. f x = pivot], [x \<leftarrow> xs. pivot < f x])"  | 
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2832  | 
|
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2833  | 
lemma part_code [code]:  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2834  | 
"part f pivot [] = ([], [], [])"  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2835  | 
"part f pivot (x # xs) = (let (lts, eqs, gts) = part f pivot xs; x' = f x in  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2836  | 
if x' < pivot then (x # lts, eqs, gts)  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2837  | 
else if x' > pivot then (lts, eqs, x # gts)  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2838  | 
else (lts, x # eqs, gts))"  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2839  | 
by (auto simp add: part_def Let_def split_def)  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2840  | 
|
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2841  | 
lemma sort_key_by_quicksort_code [code]:  | 
| 60606 | 2842  | 
"sort_key f xs =  | 
2843  | 
(case xs of  | 
|
2844  | 
[] \<Rightarrow> []  | 
|
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2845  | 
| [x] \<Rightarrow> xs  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2846  | 
| [x, y] \<Rightarrow> (if f x \<le> f y then xs else [y, x])  | 
| 60606 | 2847  | 
| _ \<Rightarrow>  | 
2848  | 
let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs  | 
|
2849  | 
in sort_key f lts @ eqs @ sort_key f gts)"  | 
|
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2850  | 
proof (cases xs)  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2851  | 
case Nil then show ?thesis by simp  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2852  | 
next  | 
| 46921 | 2853  | 
case (Cons _ ys) note hyps = Cons show ?thesis  | 
2854  | 
proof (cases ys)  | 
|
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2855  | 
case Nil with hyps show ?thesis by simp  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2856  | 
next  | 
| 46921 | 2857  | 
case (Cons _ zs) note hyps = hyps Cons show ?thesis  | 
2858  | 
proof (cases zs)  | 
|
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2859  | 
case Nil with hyps show ?thesis by auto  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2860  | 
next  | 
| 58425 | 2861  | 
case Cons  | 
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2862  | 
from sort_key_by_quicksort [of f xs]  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2863  | 
have "sort_key f xs = (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2864  | 
in sort_key f lts @ eqs @ sort_key f gts)"  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2865  | 
by (simp only: split_def Let_def part_def fst_conv snd_conv)  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2866  | 
with hyps Cons show ?thesis by (simp only: list.cases)  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2867  | 
qed  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2868  | 
qed  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2869  | 
qed  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2870  | 
|
| 
39533
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2871  | 
end  | 
| 
 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 
haftmann 
parents: 
39314 
diff
changeset
 | 
2872  | 
|
| 
40347
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2873  | 
hide_const (open) part  | 
| 
 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 
haftmann 
parents: 
40346 
diff
changeset
 | 
2874  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
2875  | 
lemma mset_remdups_subset_eq: "mset (remdups xs) \<subseteq># mset xs"  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
2876  | 
by (induct xs) (auto intro: subset_mset.order_trans)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2877  | 
|
| 60515 | 2878  | 
lemma mset_update:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2879  | 
  "i < length ls \<Longrightarrow> mset (ls[i := v]) = add_mset v (mset ls - {#ls ! i#})"
 | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2880  | 
proof (induct ls arbitrary: i)  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2881  | 
case Nil then show ?case by simp  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2882  | 
next  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2883  | 
case (Cons x xs)  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2884  | 
show ?case  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2885  | 
proof (cases i)  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2886  | 
case 0 then show ?thesis by simp  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2887  | 
next  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2888  | 
case (Suc i')  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2889  | 
with Cons show ?thesis  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2890  | 
by (cases \<open>x = xs ! i'\<close>) auto  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2891  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2892  | 
qed  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2893  | 
|
| 60515 | 2894  | 
lemma mset_swap:  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2895  | 
"i < length ls \<Longrightarrow> j < length ls \<Longrightarrow>  | 
| 60515 | 2896  | 
mset (ls[j := ls ! i, i := ls ! j]) = mset ls"  | 
2897  | 
by (cases "i = j") (simp_all add: mset_update nth_mem_mset)  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2898  | 
|
| 
73327
 
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
 
haftmann 
parents: 
73301 
diff
changeset
 | 
2899  | 
lemma mset_eq_finite:  | 
| 73301 | 2900  | 
  \<open>finite {ys. mset ys = mset xs}\<close>
 | 
2901  | 
proof -  | 
|
2902  | 
  have \<open>{ys. mset ys = mset xs} \<subseteq> {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close>
 | 
|
2903  | 
by (auto simp add: dest: mset_eq_setD mset_eq_length)  | 
|
2904  | 
  moreover have \<open>finite {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close>
 | 
|
2905  | 
using finite_lists_length_le by blast  | 
|
2906  | 
ultimately show ?thesis  | 
|
2907  | 
by (rule finite_subset)  | 
|
2908  | 
qed  | 
|
2909  | 
||
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
2910  | 
|
| 60500 | 2911  | 
subsection \<open>The multiset order\<close>  | 
2912  | 
||
| 60606 | 2913  | 
definition mult1 :: "('a \<times> 'a) set \<Rightarrow> ('a multiset \<times> 'a multiset) set" where
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2914  | 
  "mult1 r = {(N, M). \<exists>a M0 K. M = add_mset a M0 \<and> N = M0 + K \<and>
 | 
| 60607 | 2915  | 
(\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r)}"  | 
| 60606 | 2916  | 
|
2917  | 
definition mult :: "('a \<times> 'a) set \<Rightarrow> ('a multiset \<times> 'a multiset) set" where
 | 
|
| 37765 | 2918  | 
"mult r = (mult1 r)\<^sup>+"  | 
| 10249 | 2919  | 
|
| 
74858
 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 
desharna 
parents: 
74806 
diff
changeset
 | 
2920  | 
definition multp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
 | 
| 
 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 
desharna 
parents: 
74806 
diff
changeset
 | 
2921  | 
  "multp r M N \<longleftrightarrow> (M, N) \<in> mult {(x, y). r x y}"
 | 
| 
 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 
desharna 
parents: 
74806 
diff
changeset
 | 
2922  | 
|
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
2923  | 
declare multp_def[pred_set_conv]  | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
2924  | 
|
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2925  | 
lemma mult1I:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2926  | 
assumes "M = add_mset a M0" and "N = M0 + K" and "\<And>b. b \<in># K \<Longrightarrow> (b, a) \<in> r"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2927  | 
shows "(N, M) \<in> mult1 r"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2928  | 
using assms unfolding mult1_def by blast  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2929  | 
|
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2930  | 
lemma mult1E:  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2931  | 
assumes "(N, M) \<in> mult1 r"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2932  | 
obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "\<And>b. b \<in># K \<Longrightarrow> (b, a) \<in> r"  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2933  | 
using assms unfolding mult1_def by blast  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
2934  | 
|
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
2935  | 
lemma mono_mult1:  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
2936  | 
assumes "r \<subseteq> r'" shows "mult1 r \<subseteq> mult1 r'"  | 
| 
74858
 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 
desharna 
parents: 
74806 
diff
changeset
 | 
2937  | 
unfolding mult1_def using assms by blast  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
2938  | 
|
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
2939  | 
lemma mono_mult:  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
2940  | 
assumes "r \<subseteq> r'" shows "mult r \<subseteq> mult r'"  | 
| 
74858
 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 
desharna 
parents: 
74806 
diff
changeset
 | 
2941  | 
unfolding mult_def using mono_mult1[OF assms] trancl_mono by blast  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
2942  | 
|
| 74859 | 2943  | 
lemma mono_multp[mono]: "r \<le> r' \<Longrightarrow> multp r \<le> multp r'"  | 
2944  | 
unfolding le_fun_def le_bool_def  | 
|
2945  | 
proof (intro allI impI)  | 
|
2946  | 
fix M N :: "'a multiset"  | 
|
2947  | 
assume "\<forall>x xa. r x xa \<longrightarrow> r' x xa"  | 
|
2948  | 
  hence "{(x, y). r x y} \<subseteq> {(x, y). r' x y}"
 | 
|
2949  | 
by blast  | 
|
2950  | 
thus "multp r M N \<Longrightarrow> multp r' M N"  | 
|
2951  | 
unfolding multp_def  | 
|
2952  | 
by (fact mono_mult[THEN subsetD, rotated])  | 
|
2953  | 
qed  | 
|
2954  | 
||
| 23751 | 2955  | 
lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
 | 
| 
74858
 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 
desharna 
parents: 
74806 
diff
changeset
 | 
2956  | 
by (simp add: mult1_def)  | 
| 10249 | 2957  | 
|
| 74860 | 2958  | 
|
2959  | 
subsubsection \<open>Well-foundedness\<close>  | 
|
2960  | 
||
| 60608 | 2961  | 
lemma less_add:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2962  | 
assumes mult1: "(N, add_mset a M0) \<in> mult1 r"  | 
| 60608 | 2963  | 
shows  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2964  | 
"(\<exists>M. (M, M0) \<in> mult1 r \<and> N = add_mset a M) \<or>  | 
| 60608 | 2965  | 
(\<exists>K. (\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r) \<and> N = M0 + K)"  | 
2966  | 
proof -  | 
|
| 60607 | 2967  | 
let ?r = "\<lambda>K a. \<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2968  | 
let ?R = "\<lambda>N M. \<exists>a M0 K. M = add_mset a M0 \<and> N = M0 + K \<and> ?r K a"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2969  | 
obtain a' M0' K where M0: "add_mset a M0 = add_mset a' M0'"  | 
| 60608 | 2970  | 
and N: "N = M0' + K"  | 
2971  | 
and r: "?r K a'"  | 
|
2972  | 
using mult1 unfolding mult1_def by auto  | 
|
2973  | 
show ?thesis (is "?case1 \<or> ?case2")  | 
|
| 60606 | 2974  | 
proof -  | 
2975  | 
from M0 consider "M0 = M0'" "a = a'"  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
2976  | 
| K' where "M0 = add_mset a' K'" "M0' = add_mset a K'"  | 
| 60606 | 2977  | 
by atomize_elim (simp only: add_eq_conv_ex)  | 
| 18258 | 2978  | 
then show ?thesis  | 
| 60606 | 2979  | 
proof cases  | 
2980  | 
case 1  | 
|
| 11464 | 2981  | 
with N r have "?r K a \<and> N = M0 + K" by simp  | 
| 60606 | 2982  | 
then have ?case2 ..  | 
2983  | 
then show ?thesis ..  | 
|
| 10249 | 2984  | 
next  | 
| 60606 | 2985  | 
case 2  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
2986  | 
from N 2(2) have n: "N = add_mset a (K' + K)" by simp  | 
| 60606 | 2987  | 
with r 2(1) have "?R (K' + K) M0" by blast  | 
| 60608 | 2988  | 
with n have ?case1 by (simp add: mult1_def)  | 
| 60606 | 2989  | 
then show ?thesis ..  | 
| 10249 | 2990  | 
qed  | 
2991  | 
qed  | 
|
2992  | 
qed  | 
|
2993  | 
||
| 60608 | 2994  | 
lemma all_accessible:  | 
2995  | 
assumes "wf r"  | 
|
2996  | 
shows "\<forall>M. M \<in> Wellfounded.acc (mult1 r)"  | 
|
| 10249 | 2997  | 
proof  | 
2998  | 
let ?R = "mult1 r"  | 
|
| 54295 | 2999  | 
let ?W = "Wellfounded.acc ?R"  | 
| 10249 | 3000  | 
  {
 | 
3001  | 
fix M M0 a  | 
|
| 23751 | 3002  | 
assume M0: "M0 \<in> ?W"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3003  | 
and wf_hyp: "\<And>b. (b, a) \<in> r \<Longrightarrow> (\<forall>M \<in> ?W. add_mset b M \<in> ?W)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3004  | 
and acc_hyp: "\<forall>M. (M, M0) \<in> ?R \<longrightarrow> add_mset a M \<in> ?W"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3005  | 
have "add_mset a M0 \<in> ?W"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3006  | 
proof (rule accI [of "add_mset a M0"])  | 
| 10249 | 3007  | 
fix N  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3008  | 
assume "(N, add_mset a M0) \<in> ?R"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3009  | 
then consider M where "(M, M0) \<in> ?R" "N = add_mset a M"  | 
| 60608 | 3010  | 
| K where "\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r" "N = M0 + K"  | 
3011  | 
by atomize_elim (rule less_add)  | 
|
| 23751 | 3012  | 
then show "N \<in> ?W"  | 
| 60608 | 3013  | 
proof cases  | 
3014  | 
case 1  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3015  | 
from acc_hyp have "(M, M0) \<in> ?R \<longrightarrow> add_mset a M \<in> ?W" ..  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3016  | 
from this and \<open>(M, M0) \<in> ?R\<close> have "add_mset a M \<in> ?W" ..  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3017  | 
then show "N \<in> ?W" by (simp only: \<open>N = add_mset a M\<close>)  | 
| 10249 | 3018  | 
next  | 
| 60608 | 3019  | 
case 2  | 
3020  | 
from this(1) have "M0 + K \<in> ?W"  | 
|
| 10249 | 3021  | 
proof (induct K)  | 
| 18730 | 3022  | 
case empty  | 
| 23751 | 3023  | 
          from M0 show "M0 + {#} \<in> ?W" by simp
 | 
| 18730 | 3024  | 
next  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3025  | 
case (add x K)  | 
| 23751 | 3026  | 
from add.prems have "(x, a) \<in> r" by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3027  | 
with wf_hyp have "\<forall>M \<in> ?W. add_mset x M \<in> ?W" by blast  | 
| 23751 | 3028  | 
moreover from add have "M0 + K \<in> ?W" by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3029  | 
ultimately have "add_mset x (M0 + K) \<in> ?W" ..  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3030  | 
then show "M0 + (add_mset x K) \<in> ?W" by simp  | 
| 10249 | 3031  | 
qed  | 
| 60608 | 3032  | 
then show "N \<in> ?W" by (simp only: 2(2))  | 
| 10249 | 3033  | 
qed  | 
3034  | 
qed  | 
|
3035  | 
} note tedious_reasoning = this  | 
|
3036  | 
||
| 60608 | 3037  | 
show "M \<in> ?W" for M  | 
| 10249 | 3038  | 
proof (induct M)  | 
| 23751 | 3039  | 
    show "{#} \<in> ?W"
 | 
| 10249 | 3040  | 
proof (rule accI)  | 
| 23751 | 3041  | 
      fix b assume "(b, {#}) \<in> ?R"
 | 
3042  | 
with not_less_empty show "b \<in> ?W" by contradiction  | 
|
| 10249 | 3043  | 
qed  | 
3044  | 
||
| 23751 | 3045  | 
fix M a assume "M \<in> ?W"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3046  | 
from \<open>wf r\<close> have "\<forall>M \<in> ?W. add_mset a M \<in> ?W"  | 
| 10249 | 3047  | 
proof induct  | 
3048  | 
fix a  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3049  | 
assume r: "\<And>b. (b, a) \<in> r \<Longrightarrow> (\<forall>M \<in> ?W. add_mset b M \<in> ?W)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3050  | 
show "\<forall>M \<in> ?W. add_mset a M \<in> ?W"  | 
| 10249 | 3051  | 
proof  | 
| 23751 | 3052  | 
fix M assume "M \<in> ?W"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3053  | 
then show "add_mset a M \<in> ?W"  | 
| 23373 | 3054  | 
by (rule acc_induct) (rule tedious_reasoning [OF _ r])  | 
| 10249 | 3055  | 
qed  | 
3056  | 
qed  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3057  | 
from this and \<open>M \<in> ?W\<close> show "add_mset a M \<in> ?W" ..  | 
| 10249 | 3058  | 
qed  | 
3059  | 
qed  | 
|
3060  | 
||
| 74860 | 3061  | 
lemma wf_mult1: "wf r \<Longrightarrow> wf (mult1 r)"  | 
3062  | 
by (rule acc_wfI) (rule all_accessible)  | 
|
3063  | 
||
3064  | 
lemma wf_mult: "wf r \<Longrightarrow> wf (mult r)"  | 
|
3065  | 
unfolding mult_def by (rule wf_trancl) (rule wf_mult1)  | 
|
3066  | 
||
3067  | 
lemma wfP_multp: "wfP r \<Longrightarrow> wfP (multp r)"  | 
|
3068  | 
unfolding multp_def wfP_def  | 
|
3069  | 
by (simp add: wf_mult)  | 
|
| 10249 | 3070  | 
|
3071  | 
||
| 60500 | 3072  | 
subsubsection \<open>Closure-free presentation\<close>  | 
3073  | 
||
3074  | 
text \<open>One direction.\<close>  | 
|
| 10249 | 3075  | 
lemma mult_implies_one_step:  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3076  | 
assumes  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3077  | 
trans: "trans r" and  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3078  | 
MN: "(M, N) \<in> mult r"  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3079  | 
  shows "\<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and> (\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> r)"
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3080  | 
using MN unfolding mult_def mult1_def  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3081  | 
proof (induction rule: converse_trancl_induct)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3082  | 
case (base y)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3083  | 
then show ?case by force  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3084  | 
next  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3085  | 
case (step y z) note yz = this(1) and zN = this(2) and N_decomp = this(3)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3086  | 
obtain I J K where  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3087  | 
    N: "N = I + J" "z = I + K" "J \<noteq> {#}" "\<forall>k\<in>#K. \<exists>j\<in>#J. (k, j) \<in> r"
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3088  | 
using N_decomp by blast  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3089  | 
obtain a M0 K' where  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3090  | 
z: "z = add_mset a M0" and y: "y = M0 + K'" and K: "\<forall>b. b \<in># K' \<longrightarrow> (b, a) \<in> r"  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3091  | 
using yz by blast  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3092  | 
show ?case  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3093  | 
proof (cases "a \<in># K")  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3094  | 
case True  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3095  | 
moreover have "\<exists>j\<in>#J. (k, j) \<in> r" if "k \<in># K'" for k  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3096  | 
using K N trans True by (meson that transE)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3097  | 
ultimately show ?thesis  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3098  | 
      by (rule_tac x = I in exI, rule_tac x = J in exI, rule_tac x = "(K - {#a#}) + K'" in exI)
 | 
| 
64017
 
6e7bf7678518
more multiset simp rules
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63924 
diff
changeset
 | 
3099  | 
(use z y N in \<open>auto simp del: subset_mset.add_diff_assoc2 dest: in_diffD\<close>)  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3100  | 
next  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3101  | 
case False  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3102  | 
then have "a \<in># I" by (metis N(2) union_iff union_single_eq_member z)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3103  | 
    moreover have "M0 = I + K - {#a#}"
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3104  | 
using N(2) z by force  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3105  | 
ultimately show ?thesis  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3106  | 
      by (rule_tac x = "I - {#a#}" in exI, rule_tac x = "add_mset a J" in exI,
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3107  | 
rule_tac x = "K + K'" in exI)  | 
| 
64017
 
6e7bf7678518
more multiset simp rules
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63924 
diff
changeset
 | 
3108  | 
(use z y N False K in \<open>auto simp: add.assoc\<close>)  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3109  | 
qed  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3110  | 
qed  | 
| 10249 | 3111  | 
|
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3112  | 
lemma multp_implies_one_step:  | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3113  | 
  "transp R \<Longrightarrow> multp R M N \<Longrightarrow> \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and> (\<forall>k\<in>#K. \<exists>x\<in>#J. R k x)"
 | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3114  | 
by (rule mult_implies_one_step[to_pred])  | 
| 
74861
 
74ac414618e2
added lemmas multp_implies_one_step, one_step_implies_multp, and subset_implies_multp
 
desharna 
parents: 
74860 
diff
changeset
 | 
3115  | 
|
| 17161 | 3116  | 
lemma one_step_implies_mult:  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3117  | 
assumes  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3118  | 
    "J \<noteq> {#}" and
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3119  | 
"\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> r"  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3120  | 
shows "(I + K, I + J) \<in> mult r"  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3121  | 
using assms  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3122  | 
proof (induction "size J" arbitrary: I J K)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3123  | 
case 0  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3124  | 
then show ?case by auto  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3125  | 
next  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3126  | 
case (Suc n) note IH = this(1) and size_J = this(2)[THEN sym]  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3127  | 
obtain J' a where J: "J = add_mset a J'"  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3128  | 
using size_J by (blast dest: size_eq_Suc_imp_eq_union)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3129  | 
show ?case  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3130  | 
  proof (cases "J' = {#}")
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3131  | 
case True  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3132  | 
then show ?thesis  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3133  | 
using J Suc by (fastforce simp add: mult_def mult1_def)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3134  | 
next  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3135  | 
case [simp]: False  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3136  | 
    have K: "K = {#x \<in># K. (x, a) \<in> r#} + {#x \<in># K. (x, a) \<notin> r#}"
 | 
| 68992 | 3137  | 
by simp  | 
| 
63795
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3138  | 
    have "(I + K, (I + {# x \<in># K. (x, a) \<in> r #}) + J') \<in> mult r"
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3139  | 
      using IH[of J' "{# x \<in># K. (x, a) \<notin> r#}" "I + {# x \<in># K. (x, a) \<in> r#}"]
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3140  | 
J Suc.prems K size_J by (auto simp: ac_simps)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3141  | 
    moreover have "(I + {#x \<in># K. (x, a) \<in> r#} + J', I + J) \<in> mult r"
 | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3142  | 
by (fastforce simp: J mult1_def mult_def)  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3143  | 
ultimately show ?thesis  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3144  | 
unfolding mult_def by simp  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3145  | 
qed  | 
| 
 
7f6128adfe67
tuning multisets; more interpretations
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63794 
diff
changeset
 | 
3146  | 
qed  | 
| 10249 | 3147  | 
|
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3148  | 
lemma one_step_implies_multp:  | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3149  | 
  "J \<noteq> {#} \<Longrightarrow> \<forall>k\<in>#K. \<exists>j\<in>#J. R k j \<Longrightarrow> multp R (I + K) (I + J)"
 | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3150  | 
  by (rule one_step_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def, simplified])
 | 
| 
74861
 
74ac414618e2
added lemmas multp_implies_one_step, one_step_implies_multp, and subset_implies_multp
 
desharna 
parents: 
74860 
diff
changeset
 | 
3151  | 
|
| 65047 | 3152  | 
lemma subset_implies_mult:  | 
3153  | 
assumes sub: "A \<subset># B"  | 
|
3154  | 
shows "(A, B) \<in> mult r"  | 
|
3155  | 
proof -  | 
|
3156  | 
have ApBmA: "A + (B - A) = B"  | 
|
3157  | 
using sub by simp  | 
|
3158  | 
  have BmA: "B - A \<noteq> {#}"
 | 
|
3159  | 
using sub by (simp add: Diff_eq_empty_iff_mset subset_mset.less_le_not_le)  | 
|
3160  | 
thus ?thesis  | 
|
3161  | 
    by (rule one_step_implies_mult[of "B - A" "{#}" _ A, unfolded ApBmA, simplified])
 | 
|
3162  | 
qed  | 
|
3163  | 
||
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3164  | 
lemma subset_implies_multp: "A \<subset># B \<Longrightarrow> multp r A B"  | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3165  | 
  by (rule subset_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def])
 | 
| 
74861
 
74ac414618e2
added lemmas multp_implies_one_step, one_step_implies_multp, and subset_implies_multp
 
desharna 
parents: 
74860 
diff
changeset
 | 
3166  | 
|
| 77688 | 3167  | 
lemma multp_repeat_mset_repeat_msetI:  | 
3168  | 
assumes "transp R" and "multp R A B" and "n \<noteq> 0"  | 
|
3169  | 
shows "multp R (repeat_mset n A) (repeat_mset n B)"  | 
|
3170  | 
proof -  | 
|
3171  | 
from \<open>transp R\<close> \<open>multp R A B\<close> obtain I J K where  | 
|
3172  | 
    "B = I + J" and "A = I + K" and "J \<noteq> {#}" and "\<forall>k \<in># K. \<exists>x \<in># J. R k x"
 | 
|
3173  | 
by (auto dest: multp_implies_one_step)  | 
|
3174  | 
||
3175  | 
have repeat_n_A_eq: "repeat_mset n A = repeat_mset n I + repeat_mset n K"  | 
|
3176  | 
using \<open>A = I + K\<close> by simp  | 
|
3177  | 
||
3178  | 
have repeat_n_B_eq: "repeat_mset n B = repeat_mset n I + repeat_mset n J"  | 
|
3179  | 
using \<open>B = I + J\<close> by simp  | 
|
3180  | 
||
3181  | 
show ?thesis  | 
|
3182  | 
unfolding repeat_n_A_eq repeat_n_B_eq  | 
|
3183  | 
proof (rule one_step_implies_multp)  | 
|
3184  | 
    from \<open>n \<noteq> 0\<close> show "repeat_mset n J \<noteq> {#}"
 | 
|
3185  | 
      using \<open>J \<noteq> {#}\<close>
 | 
|
3186  | 
by (simp add: repeat_mset_eq_empty_iff)  | 
|
3187  | 
next  | 
|
3188  | 
show "\<forall>k \<in># repeat_mset n K. \<exists>j \<in># repeat_mset n J. R k j"  | 
|
3189  | 
using \<open>\<forall>k \<in># K. \<exists>x \<in># J. R k x\<close>  | 
|
3190  | 
by (metis count_greater_zero_iff nat_0_less_mult_iff repeat_mset.rep_eq)  | 
|
3191  | 
qed  | 
|
3192  | 
qed  | 
|
3193  | 
||
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3194  | 
|
| 
75560
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3195  | 
subsubsection \<open>Monotonicity\<close>  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3196  | 
|
| 
76401
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3197  | 
lemma multp_mono_strong:  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3198  | 
assumes "multp R M1 M2" and "transp R" and  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3199  | 
S_if_R: "\<And>x y. x \<in> set_mset M1 \<Longrightarrow> y \<in> set_mset M2 \<Longrightarrow> R x y \<Longrightarrow> S x y"  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3200  | 
shows "multp S M1 M2"  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3201  | 
proof -  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3202  | 
  obtain I J K where "M2 = I + J" and "M1 = I + K" and "J \<noteq> {#}" and "\<forall>k\<in>#K. \<exists>x\<in>#J. R k x"
 | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3203  | 
using multp_implies_one_step[OF \<open>transp R\<close> \<open>multp R M1 M2\<close>] by auto  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3204  | 
show ?thesis  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3205  | 
unfolding \<open>M2 = I + J\<close> \<open>M1 = I + K\<close>  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3206  | 
  proof (rule one_step_implies_multp[OF \<open>J \<noteq> {#}\<close>])
 | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3207  | 
show "\<forall>k\<in>#K. \<exists>j\<in>#J. S k j"  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3208  | 
using S_if_R  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3209  | 
by (metis \<open>M1 = I + K\<close> \<open>M2 = I + J\<close> \<open>\<forall>k\<in>#K. \<exists>x\<in>#J. R k x\<close> union_iff)  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3210  | 
qed  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3211  | 
qed  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3212  | 
|
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3213  | 
lemma mult_mono_strong:  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3214  | 
assumes "(M1, M2) \<in> mult r" and "trans r" and  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3215  | 
S_if_R: "\<And>x y. x \<in> set_mset M1 \<Longrightarrow> y \<in> set_mset M2 \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s"  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3216  | 
shows "(M1, M2) \<in> mult s"  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3217  | 
using assms multp_mono_strong[of "\<lambda>x y. (x, y) \<in> r" M1 M2 "\<lambda>x y. (x, y) \<in> s",  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3218  | 
unfolded multp_def transp_trans_eq, simplified]  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3219  | 
by blast  | 
| 
 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 
desharna 
parents: 
76359 
diff
changeset
 | 
3220  | 
|
| 
75584
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3221  | 
lemma monotone_on_multp_multp_image_mset:  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3222  | 
assumes "monotone_on A orda ordb f" and "transp orda"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3223  | 
  shows "monotone_on {M. set_mset M \<subseteq> A} (multp orda) (multp ordb) (image_mset f)"
 | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3224  | 
proof (rule monotone_onI)  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3225  | 
fix M1 M2  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3226  | 
assume  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3227  | 
    M1_in: "M1 \<in> {M. set_mset M \<subseteq> A}" and
 | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3228  | 
    M2_in: "M2 \<in> {M. set_mset M \<subseteq> A}" and
 | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3229  | 
M1_lt_M2: "multp orda M1 M2"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3230  | 
|
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3231  | 
from multp_implies_one_step[OF \<open>transp orda\<close> M1_lt_M2] obtain I J K where  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3232  | 
M2_eq: "M2 = I + J" and  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3233  | 
M1_eq: "M1 = I + K" and  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3234  | 
    J_neq_mempty: "J \<noteq> {#}" and
 | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3235  | 
ball_K_less: "\<forall>k\<in>#K. \<exists>x\<in>#J. orda k x"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3236  | 
by metis  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3237  | 
|
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3238  | 
have "multp ordb (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3239  | 
proof (intro one_step_implies_multp ballI)  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3240  | 
    show "image_mset f J \<noteq> {#}"
 | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3241  | 
using J_neq_mempty by simp  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3242  | 
next  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3243  | 
fix k' assume "k'\<in>#image_mset f K"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3244  | 
then obtain k where "k' = f k" and k_in: "k \<in># K"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3245  | 
by auto  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3246  | 
then obtain j where j_in: "j\<in>#J" and "orda k j"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3247  | 
using ball_K_less by auto  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3248  | 
|
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3249  | 
have "ordb (f k) (f j)"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3250  | 
proof (rule \<open>monotone_on A orda ordb f\<close>[THEN monotone_onD, OF _ _ \<open>orda k j\<close>])  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3251  | 
show "k \<in> A"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3252  | 
using M1_eq M1_in k_in by auto  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3253  | 
next  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3254  | 
show "j \<in> A"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3255  | 
using M2_eq M2_in j_in by auto  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3256  | 
qed  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3257  | 
thus "\<exists>j\<in>#image_mset f J. ordb k' j"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3258  | 
using \<open>j \<in># J\<close> \<open>k' = f k\<close> by auto  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3259  | 
qed  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3260  | 
thus "multp ordb (image_mset f M1) (image_mset f M2)"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3261  | 
by (simp add: M1_eq M2_eq)  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3262  | 
qed  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3263  | 
|
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3264  | 
lemma monotone_multp_multp_image_mset:  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3265  | 
assumes "monotone orda ordb f" and "transp orda"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3266  | 
shows "monotone (multp orda) (multp ordb) (image_mset f)"  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3267  | 
by (rule monotone_on_multp_multp_image_mset[OF assms, simplified])  | 
| 
 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 
desharna 
parents: 
75560 
diff
changeset
 | 
3268  | 
|
| 77832 | 3269  | 
lemma multp_image_mset_image_msetI:  | 
3270  | 
assumes "multp (\<lambda>x y. R (f x) (f y)) M1 M2" and "transp R"  | 
|
3271  | 
shows "multp R (image_mset f M1) (image_mset f M2)"  | 
|
3272  | 
proof -  | 
|
3273  | 
from \<open>transp R\<close> have "transp (\<lambda>x y. R (f x) (f y))"  | 
|
3274  | 
by (auto intro: transpI dest: transpD)  | 
|
3275  | 
with \<open>multp (\<lambda>x y. R (f x) (f y)) M1 M2\<close> obtain I J K where  | 
|
3276  | 
    "M2 = I + J" and "M1 = I + K" and "J \<noteq> {#}" and "\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)"
 | 
|
3277  | 
using multp_implies_one_step by blast  | 
|
3278  | 
||
3279  | 
have "multp R (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)"  | 
|
3280  | 
proof (rule one_step_implies_multp)  | 
|
3281  | 
    show "image_mset f J \<noteq> {#}"
 | 
|
3282  | 
      by (simp add: \<open>J \<noteq> {#}\<close>)
 | 
|
3283  | 
next  | 
|
3284  | 
show "\<forall>k\<in>#image_mset f K. \<exists>j\<in>#image_mset f J. R k j"  | 
|
3285  | 
by (simp add: \<open>\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)\<close>)  | 
|
3286  | 
qed  | 
|
3287  | 
thus ?thesis  | 
|
3288  | 
by (simp add: \<open>M1 = I + K\<close> \<open>M2 = I + J\<close>)  | 
|
3289  | 
qed  | 
|
3290  | 
||
| 
75560
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3291  | 
lemma multp_image_mset_image_msetD:  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3292  | 
assumes  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3293  | 
"multp R (image_mset f A) (image_mset f B)" and  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3294  | 
"transp R" and  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3295  | 
inj_on_f: "inj_on f (set_mset A \<union> set_mset B)"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3296  | 
shows "multp (\<lambda>x y. R (f x) (f y)) A B"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3297  | 
proof -  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3298  | 
from assms(1,2) obtain I J K where  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3299  | 
f_B_eq: "image_mset f B = I + J" and  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3300  | 
f_A_eq: "image_mset f A = I + K" and  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3301  | 
    J_neq_mempty: "J \<noteq> {#}" and
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3302  | 
ball_K_less: "\<forall>k\<in>#K. \<exists>x\<in>#J. R k x"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3303  | 
by (auto dest: multp_implies_one_step)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3304  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3305  | 
from f_B_eq obtain I' J' where  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3306  | 
B_def: "B = I' + J'" and I_def: "I = image_mset f I'" and J_def: "J = image_mset f J'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3307  | 
using image_mset_eq_plusD by blast  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3308  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3309  | 
from inj_on_f have inj_on_f': "inj_on f (set_mset A \<union> set_mset I')"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3310  | 
by (rule inj_on_subset) (auto simp add: B_def)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3311  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3312  | 
from f_A_eq obtain K' where  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3313  | 
A_def: "A = I' + K'" and K_def: "K = image_mset f K'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3314  | 
by (auto simp: I_def dest: image_mset_eq_image_mset_plusD[OF _ inj_on_f'])  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3315  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3316  | 
show ?thesis  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3317  | 
unfolding A_def B_def  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3318  | 
proof (intro one_step_implies_multp ballI)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3319  | 
    from J_neq_mempty show "J' \<noteq> {#}"
 | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3320  | 
by (simp add: J_def)  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3321  | 
next  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3322  | 
fix k assume "k \<in># K'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3323  | 
with ball_K_less obtain j' where "j' \<in># J" and "R (f k) j'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3324  | 
using K_def by auto  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3325  | 
moreover then obtain j where "j \<in># J'" and "f j = j'"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3326  | 
using J_def by auto  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3327  | 
ultimately show "\<exists>j\<in>#J'. R (f k) (f j)"  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3328  | 
by blast  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3329  | 
qed  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3330  | 
qed  | 
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3331  | 
|
| 
 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 
desharna 
parents: 
75467 
diff
changeset
 | 
3332  | 
|
| 
74862
 
aa51e974b688
added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
 
desharna 
parents: 
74861 
diff
changeset
 | 
3333  | 
subsubsection \<open>The multiset extension is cancellative for multiset union\<close>  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3334  | 
|
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3335  | 
lemma mult_cancel:  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3336  | 
assumes "trans s" and "irrefl_on (set_mset Z) s"  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3337  | 
shows "(X + Z, Y + Z) \<in> mult s \<longleftrightarrow> (X, Y) \<in> mult s" (is "?L \<longleftrightarrow> ?R")  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3338  | 
proof  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3339  | 
assume ?L thus ?R  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3340  | 
using \<open>irrefl_on (set_mset Z) s\<close>  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3341  | 
proof (induct Z)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3342  | 
case (add z Z)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3343  | 
    obtain X' Y' Z' where *: "add_mset z X + Z = Z' + X'" "add_mset z Y + Z = Z' + Y'" "Y' \<noteq> {#}"
 | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3344  | 
"\<forall>x \<in> set_mset X'. \<exists>y \<in> set_mset Y'. (x, y) \<in> s"  | 
| 64911 | 3345  | 
using mult_implies_one_step[OF \<open>trans s\<close> add(2)] by auto  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3346  | 
consider Z2 where "Z' = add_mset z Z2" | X2 Y2 where "X' = add_mset z X2" "Y' = add_mset z Y2"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3347  | 
using *(1,2) by (metis add_mset_remove_trivial_If insert_iff set_mset_add_mset_insert union_iff)  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3348  | 
thus ?case  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3349  | 
proof (cases)  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3350  | 
case 1 thus ?thesis  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3351  | 
using * one_step_implies_mult[of Y' X' s Z2] add(3)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3352  | 
        by (auto simp: add.commute[of _ "{#_#}"] add.assoc intro: add(1) elim: irrefl_on_subset)
 | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3353  | 
next  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3354  | 
case 2 then obtain y where "y \<in> set_mset Y2" "(z, y) \<in> s"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3355  | 
using *(4) \<open>irrefl_on (set_mset (add_mset z Z)) s\<close>  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3356  | 
by (auto simp: irrefl_on_def)  | 
| 64911 | 3357  | 
moreover from this transD[OF \<open>trans s\<close> _ this(2)]  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3358  | 
have "x' \<in> set_mset X2 \<Longrightarrow> \<exists>y \<in> set_mset Y2. (x', y) \<in> s" for x'  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3359  | 
using 2 *(4)[rule_format, of x'] by auto  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3360  | 
ultimately show ?thesis  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3361  | 
using * one_step_implies_mult[of Y2 X2 s Z'] 2 add(3)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3362  | 
        by (force simp: add.commute[of "{#_#}"] add.assoc[symmetric] intro: add(1)
 | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3363  | 
elim: irrefl_on_subset)  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3364  | 
qed  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3365  | 
qed auto  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3366  | 
next  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3367  | 
assume ?R then obtain I J K  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3368  | 
    where "Y = I + J" "X = I + K" "J \<noteq> {#}" "\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> s"
 | 
| 64911 | 3369  | 
using mult_implies_one_step[OF \<open>trans s\<close>] by blast  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3370  | 
thus ?L using one_step_implies_mult[of J K s "I + Z"] by (auto simp: ac_simps)  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3371  | 
qed  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3372  | 
|
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3373  | 
lemma multp_cancel:  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3374  | 
"transp R \<Longrightarrow> irreflp_on (set_mset Z) R \<Longrightarrow> multp R (X + Z) (Y + Z) \<longleftrightarrow> multp R X Y"  | 
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3375  | 
by (rule mult_cancel[to_pred])  | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3376  | 
|
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3377  | 
lemma mult_cancel_add_mset:  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3378  | 
  "trans r \<Longrightarrow> irrefl_on {z} r \<Longrightarrow>
 | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3379  | 
((add_mset z X, add_mset z Y) \<in> mult r) = ((X, Y) \<in> mult r)"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3380  | 
  by (rule mult_cancel[of _ "{#_#}", simplified])
 | 
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3381  | 
|
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3382  | 
lemma multp_cancel_add_mset:  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3383  | 
  "transp R \<Longrightarrow> irreflp_on {z} R \<Longrightarrow> multp R (add_mset z X) (add_mset z Y) = multp R X Y"
 | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3384  | 
by (rule mult_cancel_add_mset[to_pred, folded bot_set_def])  | 
| 
74862
 
aa51e974b688
added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
 
desharna 
parents: 
74861 
diff
changeset
 | 
3385  | 
|
| 
74804
 
5749fefd3fa0
simplified mult_cancel_max and introduced orginal lemma as mult_cancel_max0
 
desharna 
parents: 
74803 
diff
changeset
 | 
3386  | 
lemma mult_cancel_max0:  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3387  | 
assumes "trans s" and "irrefl_on (set_mset X \<inter> set_mset Y) s"  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3388  | 
shows "(X, Y) \<in> mult s \<longleftrightarrow> (X - X \<inter># Y, Y - X \<inter># Y) \<in> mult s" (is "?L \<longleftrightarrow> ?R")  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3389  | 
proof -  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3390  | 
have "(X - X \<inter># Y + X \<inter># Y, Y - X \<inter># Y + X \<inter># Y) \<in> mult s \<longleftrightarrow> (X - X \<inter># Y, Y - X \<inter># Y) \<in> mult s"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3391  | 
proof (rule mult_cancel)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3392  | 
from assms show "trans s"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3393  | 
by simp  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3394  | 
next  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3395  | 
from assms show "irrefl_on (set_mset (X \<inter># Y)) s"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3396  | 
by simp  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3397  | 
qed  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3398  | 
moreover have "X - X \<inter># Y + X \<inter># Y = X" "Y - X \<inter># Y + X \<inter># Y = Y"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3399  | 
by (auto simp flip: count_inject)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3400  | 
ultimately show ?thesis  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3401  | 
by simp  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3402  | 
qed  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3403  | 
|
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3404  | 
lemma mult_cancel_max:  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3405  | 
"trans r \<Longrightarrow> irrefl_on (set_mset X \<inter> set_mset Y) r \<Longrightarrow>  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3406  | 
(X, Y) \<in> mult r \<longleftrightarrow> (X - Y, Y - X) \<in> mult r"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3407  | 
by (rule mult_cancel_max0[simplified])  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3408  | 
|
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3409  | 
lemma multp_cancel_max:  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3410  | 
"transp R \<Longrightarrow> irreflp_on (set_mset X \<inter> set_mset Y) R \<Longrightarrow> multp R X Y \<longleftrightarrow> multp R (X - Y) (Y - X)"  | 
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3411  | 
by (rule mult_cancel_max[to_pred])  | 
| 
74862
 
aa51e974b688
added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
 
desharna 
parents: 
74861 
diff
changeset
 | 
3412  | 
|
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3413  | 
|
| 
77049
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3414  | 
subsubsection \<open>Strict partial-order properties\<close>  | 
| 74864 | 3415  | 
|
3416  | 
lemma mult1_lessE:  | 
|
3417  | 
  assumes "(N, M) \<in> mult1 {(a, b). r a b}" and "asymp r"
 | 
|
3418  | 
obtains a M0 K where "M = add_mset a M0" "N = M0 + K"  | 
|
3419  | 
"a \<notin># K" "\<And>b. b \<in># K \<Longrightarrow> r b a"  | 
|
3420  | 
proof -  | 
|
3421  | 
from assms obtain a M0 K where "M = add_mset a M0" "N = M0 + K" and  | 
|
3422  | 
*: "b \<in># K \<Longrightarrow> r b a" for b by (blast elim: mult1E)  | 
|
3423  | 
moreover from * [of a] have "a \<notin># K"  | 
|
| 
76682
 
e260dabc88e6
added predicates asym_on and asymp_on and redefined asym and asymp to be abbreviations
 
desharna 
parents: 
76611 
diff
changeset
 | 
3424  | 
using \<open>asymp r\<close> by (meson asympD)  | 
| 74864 | 3425  | 
ultimately show thesis by (auto intro: that)  | 
3426  | 
qed  | 
|
3427  | 
||
| 
74865
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3428  | 
lemma trans_mult: "trans r \<Longrightarrow> trans (mult r)"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3429  | 
by (simp add: mult_def)  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3430  | 
|
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3431  | 
lemma transp_multp: "transp r \<Longrightarrow> transp (multp r)"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3432  | 
unfolding multp_def transp_trans_eq  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3433  | 
  by (fact trans_mult[of "{(x, y). r x y}" for r, folded transp_trans])
 | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3434  | 
|
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3435  | 
lemma irrefl_mult:  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3436  | 
assumes "trans r" "irrefl r"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3437  | 
shows "irrefl (mult r)"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3438  | 
proof (intro irreflI notI)  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3439  | 
fix M  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3440  | 
assume "(M, M) \<in> mult r"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3441  | 
then obtain I J K where "M = I + J" and "M = I + K"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3442  | 
    and "J \<noteq> {#}" and "(\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset J. (k, j) \<in> r)"
 | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3443  | 
using mult_implies_one_step[OF \<open>trans r\<close>] by blast  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3444  | 
  then have *: "K \<noteq> {#}" and **: "\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset K. (k, j) \<in> r" by auto
 | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3445  | 
have "finite (set_mset K)" by simp  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3446  | 
  hence "set_mset K = {}"
 | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3447  | 
using **  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3448  | 
proof (induction rule: finite_induct)  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3449  | 
case empty  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3450  | 
thus ?case by simp  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3451  | 
next  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3452  | 
case (insert x F)  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3453  | 
have False  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3454  | 
using \<open>irrefl r\<close>[unfolded irrefl_def, rule_format]  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3455  | 
using \<open>trans r\<close>[THEN transD]  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3456  | 
by (metis equals0D insert.IH insert.prems insertE insertI1 insertI2)  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3457  | 
thus ?case ..  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3458  | 
qed  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3459  | 
with * show False by simp  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3460  | 
qed  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3461  | 
|
| 
76589
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3462  | 
lemma irreflp_multp: "transp R \<Longrightarrow> irreflp R \<Longrightarrow> irreflp (multp R)"  | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3463  | 
  by (rule irrefl_mult[of "{(x, y). r x y}" for r,
 | 
| 
 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 
desharna 
parents: 
76570 
diff
changeset
 | 
3464  | 
folded transp_trans_eq irreflp_irrefl_eq, simplified, folded multp_def])  | 
| 
74865
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3465  | 
|
| 74864 | 3466  | 
instantiation multiset :: (preorder) order begin  | 
3467  | 
||
3468  | 
definition less_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool"  | 
|
3469  | 
where "M < N \<longleftrightarrow> multp (<) M N"  | 
|
3470  | 
||
3471  | 
definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool"  | 
|
3472  | 
where "less_eq_multiset M N \<longleftrightarrow> M < N \<or> M = N"  | 
|
3473  | 
||
3474  | 
instance  | 
|
| 
74865
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3475  | 
proof intro_classes  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3476  | 
fix M N :: "'a multiset"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3477  | 
show "(M < N) = (M \<le> N \<and> \<not> N \<le> M)"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3478  | 
unfolding less_eq_multiset_def less_multiset_def  | 
| 
76749
 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 
desharna 
parents: 
76682 
diff
changeset
 | 
3479  | 
by (metis irreflp_def irreflp_on_less irreflp_multp transpE transp_on_less transp_multp)  | 
| 
74865
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3480  | 
next  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3481  | 
fix M :: "'a multiset"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3482  | 
show "M \<le> M"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3483  | 
unfolding less_eq_multiset_def  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3484  | 
by simp  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3485  | 
next  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3486  | 
fix M1 M2 M3 :: "'a multiset"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3487  | 
show "M1 \<le> M2 \<Longrightarrow> M2 \<le> M3 \<Longrightarrow> M1 \<le> M3"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3488  | 
unfolding less_eq_multiset_def less_multiset_def  | 
| 
76749
 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 
desharna 
parents: 
76682 
diff
changeset
 | 
3489  | 
using transp_multp[OF transp_on_less, THEN transpD]  | 
| 
74865
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3490  | 
by blast  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3491  | 
next  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3492  | 
fix M N :: "'a multiset"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3493  | 
show "M \<le> N \<Longrightarrow> N \<le> M \<Longrightarrow> M = N"  | 
| 
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3494  | 
unfolding less_eq_multiset_def less_multiset_def  | 
| 
76749
 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 
desharna 
parents: 
76682 
diff
changeset
 | 
3495  | 
using transp_multp[OF transp_on_less, THEN transpD]  | 
| 
 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 
desharna 
parents: 
76682 
diff
changeset
 | 
3496  | 
using irreflp_multp[OF transp_on_less irreflp_on_less, unfolded irreflp_def, rule_format]  | 
| 
74865
 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 
desharna 
parents: 
74864 
diff
changeset
 | 
3497  | 
by blast  | 
| 74864 | 3498  | 
qed  | 
3499  | 
||
3500  | 
end  | 
|
3501  | 
||
3502  | 
lemma mset_le_irrefl [elim!]:  | 
|
3503  | 
fixes M :: "'a::preorder multiset"  | 
|
3504  | 
shows "M < M \<Longrightarrow> R"  | 
|
3505  | 
by simp  | 
|
3506  | 
||
| 
74868
 
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
 
desharna 
parents: 
74865 
diff
changeset
 | 
3507  | 
lemma wfP_less_multiset[simp]:  | 
| 
 
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
 
desharna 
parents: 
74865 
diff
changeset
 | 
3508  | 
  assumes wfP_less: "wfP ((<) :: ('a :: preorder) \<Rightarrow> 'a \<Rightarrow> bool)"
 | 
| 
 
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
 
desharna 
parents: 
74865 
diff
changeset
 | 
3509  | 
shows "wfP ((<) :: 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool)"  | 
| 
 
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
 
desharna 
parents: 
74865 
diff
changeset
 | 
3510  | 
using wfP_multp[OF wfP_less] less_multiset_def  | 
| 
 
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
 
desharna 
parents: 
74865 
diff
changeset
 | 
3511  | 
by (metis wfPUNIVI wfP_induct)  | 
| 
 
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
 
desharna 
parents: 
74865 
diff
changeset
 | 
3512  | 
|
| 74864 | 3513  | 
|
| 
77049
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3514  | 
subsubsection \<open>Strict total-order properties\<close>  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3515  | 
|
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3516  | 
lemma total_on_mult:  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3517  | 
assumes "total_on A r" and "trans r" and "\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3518  | 
shows "total_on B (mult r)"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3519  | 
proof (rule total_onI)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3520  | 
fix M1 M2 assume "M1 \<in> B" and "M2 \<in> B" and "M1 \<noteq> M2"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3521  | 
let ?I = "M1 \<inter># M2"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3522  | 
show "(M1, M2) \<in> mult r \<or> (M2, M1) \<in> mult r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3523  | 
  proof (cases "M1 - ?I = {#} \<or> M2 - ?I = {#}")
 | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3524  | 
case True  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3525  | 
with \<open>M1 \<noteq> M2\<close> show ?thesis  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3526  | 
by (metis Diff_eq_empty_iff_mset diff_intersect_left_idem diff_intersect_right_idem  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3527  | 
subset_implies_mult subset_mset.less_le)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3528  | 
next  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3529  | 
case False  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3530  | 
from assms(1) have "total_on (set_mset (M1 - ?I)) r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3531  | 
by (meson \<open>M1 \<in> B\<close> assms(3) diff_subset_eq_self set_mset_mono total_on_subset)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3532  | 
with False obtain greatest1 where  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3533  | 
greatest1_in: "greatest1 \<in># M1 - ?I" and  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3534  | 
greatest1_greatest: "\<forall>x \<in># M1 - ?I. greatest1 \<noteq> x \<longrightarrow> (x, greatest1) \<in> r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3535  | 
using Multiset.bex_greatest_element[to_set, of "M1 - ?I" r]  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3536  | 
by (metis assms(2) subset_UNIV trans_on_subset)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3537  | 
|
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3538  | 
from assms(1) have "total_on (set_mset (M2 - ?I)) r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3539  | 
by (meson \<open>M2 \<in> B\<close> assms(3) diff_subset_eq_self set_mset_mono total_on_subset)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3540  | 
with False obtain greatest2 where  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3541  | 
greatest2_in: "greatest2 \<in># M2 - ?I" and  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3542  | 
greatest2_greatest: "\<forall>x \<in># M2 - ?I. greatest2 \<noteq> x \<longrightarrow> (x, greatest2) \<in> r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3543  | 
using Multiset.bex_greatest_element[to_set, of "M2 - ?I" r]  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3544  | 
by (metis assms(2) subset_UNIV trans_on_subset)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3545  | 
|
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3546  | 
have "greatest1 \<noteq> greatest2"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3547  | 
using greatest1_in \<open>greatest2 \<in># M2 - ?I\<close>  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3548  | 
by (metis diff_intersect_left_idem diff_intersect_right_idem dual_order.eq_iff in_diff_count  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3549  | 
in_diff_countE le_add_same_cancel2 less_irrefl zero_le)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3550  | 
hence "(greatest1, greatest2) \<in> r \<or> (greatest2, greatest1) \<in> r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3551  | 
using \<open>total_on A r\<close>[unfolded total_on_def, rule_format, of greatest1 greatest2]  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3552  | 
\<open>M1 \<in> B\<close> \<open>M2 \<in> B\<close> greatest1_in greatest2_in assms(3)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3553  | 
by (meson in_diffD in_mono)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3554  | 
thus ?thesis  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3555  | 
proof (elim disjE)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3556  | 
assume "(greatest1, greatest2) \<in> r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3557  | 
have "(?I + (M1 - ?I), ?I + (M2 - ?I)) \<in> mult r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3558  | 
proof (rule one_step_implies_mult[of "M2 - ?I" "M1 - ?I" r ?I])  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3559  | 
        show "M2 - ?I \<noteq> {#}"
 | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3560  | 
using False by force  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3561  | 
next  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3562  | 
show "\<forall>k\<in>#M1 - ?I. \<exists>j\<in>#M2 - ?I. (k, j) \<in> r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3563  | 
using \<open>(greatest1, greatest2) \<in> r\<close> greatest2_in greatest1_greatest  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3564  | 
by (metis assms(2) transD)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3565  | 
qed  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3566  | 
hence "(M1, M2) \<in> mult r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3567  | 
by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3568  | 
subset_mset.inf.cobounded2)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3569  | 
thus "(M1, M2) \<in> mult r \<or> (M2, M1) \<in> mult r" ..  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3570  | 
next  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3571  | 
assume "(greatest2, greatest1) \<in> r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3572  | 
have "(?I + (M2 - ?I), ?I + (M1 - ?I)) \<in> mult r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3573  | 
proof (rule one_step_implies_mult[of "M1 - ?I" "M2 - ?I" r ?I])  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3574  | 
        show "M1 - M1 \<inter># M2 \<noteq> {#}"
 | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3575  | 
using False by force  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3576  | 
next  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3577  | 
show "\<forall>k\<in>#M2 - ?I. \<exists>j\<in>#M1 - ?I. (k, j) \<in> r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3578  | 
using \<open>(greatest2, greatest1) \<in> r\<close> greatest1_in greatest2_greatest  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3579  | 
by (metis assms(2) transD)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3580  | 
qed  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3581  | 
hence "(M2, M1) \<in> mult r"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3582  | 
by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3583  | 
subset_mset.inf.cobounded2)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3584  | 
thus "(M1, M2) \<in> mult r \<or> (M2, M1) \<in> mult r" ..  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3585  | 
qed  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3586  | 
qed  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3587  | 
qed  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3588  | 
|
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3589  | 
lemma total_mult: "total r \<Longrightarrow> trans r \<Longrightarrow> total (mult r)"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3590  | 
by (rule total_on_mult[of UNIV r UNIV, simplified])  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3591  | 
|
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3592  | 
lemma totalp_on_multp:  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3593  | 
"totalp_on A R \<Longrightarrow> transp R \<Longrightarrow> (\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A) \<Longrightarrow> totalp_on B (multp R)"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3594  | 
  using total_on_mult[of A "{(x,y). R x y}" B, to_pred]
 | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3595  | 
by (simp add: multp_def total_on_def totalp_on_def)  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3596  | 
|
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3597  | 
lemma totalp_multp: "totalp R \<Longrightarrow> transp R \<Longrightarrow> totalp (multp R)"  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3598  | 
by (rule totalp_on_multp[of UNIV R UNIV, simplified])  | 
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3599  | 
|
| 
 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 
desharna 
parents: 
76755 
diff
changeset
 | 
3600  | 
|
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3601  | 
subsection \<open>Quasi-executable version of the multiset extension\<close>  | 
| 
63088
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3602  | 
|
| 
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3603  | 
text \<open>  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3604  | 
Predicate variants of \<open>mult\<close> and the reflexive closure of \<open>mult\<close>, which are  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3605  | 
executable whenever the given predicate \<open>P\<close> is. Together with the standard  | 
| 67398 | 3606  | 
code equations for \<open>(\<inter>#\<close>) and \<open>(-\<close>) this should yield quadratic  | 
| 74803 | 3607  | 
(with respect to calls to \<open>P\<close>) implementations of \<open>multp_code\<close> and \<open>multeqp_code\<close>.  | 
| 
63088
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3608  | 
\<close>  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3609  | 
|
| 74803 | 3610  | 
definition multp_code :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
 | 
3611  | 
"multp_code P N M =  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3612  | 
(let Z = M \<inter># N; X = M - Z in  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3613  | 
    X \<noteq> {#} \<and> (let Y = N - Z in (\<forall>y \<in> set_mset Y. \<exists>x \<in> set_mset X. P y x)))"
 | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3614  | 
|
| 74803 | 3615  | 
definition multeqp_code :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
 | 
3616  | 
"multeqp_code P N M =  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3617  | 
(let Z = M \<inter># N; X = M - Z; Y = N - Z in  | 
| 
63088
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3618  | 
(\<forall>y \<in> set_mset Y. \<exists>x \<in> set_mset X. P y x))"  | 
| 
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3619  | 
|
| 74805 | 3620  | 
lemma multp_code_iff_mult:  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3621  | 
assumes "irrefl_on (set_mset N \<inter> set_mset M) R" and "trans R" and  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3622  | 
[simp]: "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R"  | 
| 74803 | 3623  | 
shows "multp_code P N M \<longleftrightarrow> (N, M) \<in> mult R" (is "?L \<longleftrightarrow> ?R")  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3624  | 
proof -  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3625  | 
have *: "M \<inter># N + (N - M \<inter># N) = N" "M \<inter># N + (M - M \<inter># N) = M"  | 
| 68406 | 3626  | 
    "(M - M \<inter># N) \<inter># (N - M \<inter># N) = {#}" by (auto simp flip: count_inject)
 | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3627  | 
show ?thesis  | 
| 
63088
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3628  | 
proof  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3629  | 
assume ?L thus ?R  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3630  | 
using one_step_implies_mult[of "M - M \<inter># N" "N - M \<inter># N" R "M \<inter># N"] *  | 
| 74803 | 3631  | 
by (auto simp: multp_code_def Let_def)  | 
| 
63088
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3632  | 
next  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3633  | 
    { fix I J K :: "'a multiset" assume "(I + J) \<inter># (I + K) = {#}"
 | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3634  | 
      then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty)
 | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3635  | 
} note [dest!] = this  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3636  | 
assume ?R thus ?L  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3637  | 
using mult_cancel_max  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3638  | 
using mult_implies_one_step[OF assms(2), of "N - M \<inter># N" "M - M \<inter># N"]  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3639  | 
mult_cancel_max[OF assms(2,1)] * by (auto simp: multp_code_def)  | 
| 
63088
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3640  | 
qed  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3641  | 
qed  | 
| 
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3642  | 
|
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3643  | 
lemma multp_code_iff_multp:  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3644  | 
"irreflp_on (set_mset M \<inter> set_mset N) R \<Longrightarrow> transp R \<Longrightarrow> multp_code R M N \<longleftrightarrow> multp R M N"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3645  | 
using multp_code_iff_mult[simplified, to_pred, of M N R R] by simp  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3646  | 
|
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3647  | 
lemma multp_code_eq_multp:  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3648  | 
assumes "irreflp R" and "transp R"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3649  | 
shows "multp_code R = multp R"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3650  | 
proof (intro ext)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3651  | 
fix M N  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3652  | 
show "multp_code R M N = multp R M N"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3653  | 
proof (rule multp_code_iff_multp)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3654  | 
from assms show "irreflp_on (set_mset M \<inter> set_mset N) R"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3655  | 
by (auto intro: irreflp_on_subset)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3656  | 
next  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3657  | 
from assms show "transp R"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3658  | 
by simp  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3659  | 
qed  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3660  | 
qed  | 
| 
74863
 
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
 
desharna 
parents: 
74862 
diff
changeset
 | 
3661  | 
|
| 74805 | 3662  | 
lemma multeqp_code_iff_reflcl_mult:  | 
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3663  | 
assumes "irrefl_on (set_mset N \<inter> set_mset M) R" and "trans R" and "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R"  | 
| 74803 | 3664  | 
shows "multeqp_code P N M \<longleftrightarrow> (N, M) \<in> (mult R)\<^sup>="  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3665  | 
proof -  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3666  | 
  { assume "N \<noteq> M" "M - M \<inter># N = {#}"
 | 
| 68406 | 3667  | 
then obtain y where "count N y \<noteq> count M y" by (auto simp flip: count_inject)  | 
| 64911 | 3668  | 
    then have "\<exists>y. count M y < count N y" using \<open>M - M \<inter># N = {#}\<close>
 | 
| 68406 | 3669  | 
by (auto simp flip: count_inject dest!: le_neq_implies_less fun_cong[of _ _ y])  | 
| 
63660
 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> 
parents: 
63560 
diff
changeset
 | 
3670  | 
}  | 
| 74803 | 3671  | 
then have "multeqp_code P N M \<longleftrightarrow> multp_code P N M \<or> N = M"  | 
3672  | 
by (auto simp: multeqp_code_def multp_code_def Let_def in_diff_count)  | 
|
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3673  | 
thus ?thesis  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3674  | 
using multp_code_iff_mult[OF assms] by simp  | 
| 
63088
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3675  | 
qed  | 
| 
 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 
haftmann 
parents: 
63060 
diff
changeset
 | 
3676  | 
|
| 
76611
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3677  | 
lemma multeqp_code_iff_reflclp_multp:  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3678  | 
"irreflp_on (set_mset M \<inter> set_mset N) R \<Longrightarrow> transp R \<Longrightarrow> multeqp_code R M N \<longleftrightarrow> (multp R)\<^sup>=\<^sup>= M N"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3679  | 
using multeqp_code_iff_reflcl_mult[simplified, to_pred, of M N R R] by simp  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3680  | 
|
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3681  | 
lemma multeqp_code_eq_reflclp_multp:  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3682  | 
assumes "irreflp R" and "transp R"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3683  | 
shows "multeqp_code R = (multp R)\<^sup>=\<^sup>="  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3684  | 
proof (intro ext)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3685  | 
fix M N  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3686  | 
show "multeqp_code R M N \<longleftrightarrow> (multp R)\<^sup>=\<^sup>= M N"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3687  | 
proof (rule multeqp_code_iff_reflclp_multp)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3688  | 
from assms show "irreflp_on (set_mset M \<inter> set_mset N) R"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3689  | 
by (auto intro: irreflp_on_subset)  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3690  | 
next  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3691  | 
from assms show "transp R"  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3692  | 
by simp  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3693  | 
qed  | 
| 
 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 
desharna 
parents: 
76589 
diff
changeset
 | 
3694  | 
qed  | 
| 
74863
 
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
 
desharna 
parents: 
74862 
diff
changeset
 | 
3695  | 
|
| 
 
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
 
desharna 
parents: 
74862 
diff
changeset
 | 
3696  | 
|
| 60500 | 3697  | 
subsubsection \<open>Monotonicity of multiset union\<close>  | 
| 10249 | 3698  | 
|
| 60606 | 3699  | 
lemma mult1_union: "(B, D) \<in> mult1 r \<Longrightarrow> (C + B, C + D) \<in> mult1 r"  | 
| 64076 | 3700  | 
by (force simp: mult1_def)  | 
| 10249 | 3701  | 
|
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3702  | 
lemma union_le_mono2: "B < D \<Longrightarrow> C + B < C + (D::'a::preorder multiset)"  | 
| 74864 | 3703  | 
apply (unfold less_multiset_def multp_def mult_def)  | 
| 26178 | 3704  | 
apply (erule trancl_induct)  | 
| 
40249
 
cd404ecb9107
Remove unnecessary premise of mult1_union
 
Lars Noschinski <noschinl@in.tum.de> 
parents: 
39533 
diff
changeset
 | 
3705  | 
apply (blast intro: mult1_union)  | 
| 
 
cd404ecb9107
Remove unnecessary premise of mult1_union
 
Lars Noschinski <noschinl@in.tum.de> 
parents: 
39533 
diff
changeset
 | 
3706  | 
apply (blast intro: mult1_union trancl_trans)  | 
| 26178 | 3707  | 
done  | 
| 10249 | 3708  | 
|
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3709  | 
lemma union_le_mono1: "B < D \<Longrightarrow> B + C < D + (C::'a::preorder multiset)"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
3710  | 
apply (subst add.commute [of B C])  | 
| 
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
3711  | 
apply (subst add.commute [of D C])  | 
| 
63388
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3712  | 
apply (erule union_le_mono2)  | 
| 26178 | 3713  | 
done  | 
| 10249 | 3714  | 
|
| 17161 | 3715  | 
lemma union_less_mono:  | 
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3716  | 
fixes A B C D :: "'a::preorder multiset"  | 
| 
63388
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3717  | 
shows "A < C \<Longrightarrow> B < D \<Longrightarrow> A + B < C + D"  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3718  | 
by (blast intro!: union_le_mono1 union_le_mono2 less_trans)  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3719  | 
|
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3720  | 
instantiation multiset :: (preorder) ordered_ab_semigroup_add  | 
| 
63388
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3721  | 
begin  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3722  | 
instance  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3723  | 
by standard (auto simp add: less_eq_multiset_def intro: union_le_mono2)  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3724  | 
end  | 
| 15072 | 3725  | 
|
| 
63409
 
3f3223b90239
moved lemmas and locales around (with minor incompatibilities)
 
blanchet 
parents: 
63388 
diff
changeset
 | 
3726  | 
|
| 60500 | 3727  | 
subsubsection \<open>Termination proofs with multiset orders\<close>  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3728  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3729  | 
lemma multi_member_skip: "x \<in># XS \<Longrightarrow> x \<in># {# y #} + XS"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3730  | 
  and multi_member_this: "x \<in># {# x #} + XS"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3731  | 
  and multi_member_last: "x \<in># {# x #}"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3732  | 
by auto  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3733  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3734  | 
definition "ms_strict = mult pair_less"  | 
| 37765 | 3735  | 
definition "ms_weak = ms_strict \<union> Id"  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3736  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3737  | 
lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3738  | 
unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3739  | 
by (auto intro: wf_mult1 wf_trancl simp: mult_def)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3740  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3741  | 
lemma smsI:  | 
| 60495 | 3742  | 
"(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z + B) \<in> ms_strict"  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3743  | 
unfolding ms_strict_def  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3744  | 
by (rule one_step_implies_mult) (auto simp add: max_strict_def pair_less_def elim!:max_ext.cases)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3745  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3746  | 
lemma wmsI:  | 
| 60495 | 3747  | 
  "(set_mset A, set_mset B) \<in> max_strict \<or> A = {#} \<and> B = {#}
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3748  | 
\<Longrightarrow> (Z + A, Z + B) \<in> ms_weak"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3749  | 
unfolding ms_weak_def ms_strict_def  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3750  | 
by (auto simp add: pair_less_def max_strict_def elim!:max_ext.cases intro: one_step_implies_mult)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3751  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3752  | 
inductive pw_leq  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3753  | 
where  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3754  | 
  pw_leq_empty: "pw_leq {#} {#}"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3755  | 
| pw_leq_step:  "\<lbrakk>(x,y) \<in> pair_leq; pw_leq X Y \<rbrakk> \<Longrightarrow> pw_leq ({#x#} + X) ({#y#} + Y)"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3756  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3757  | 
lemma pw_leq_lstep:  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3758  | 
  "(x, y) \<in> pair_leq \<Longrightarrow> pw_leq {#x#} {#y#}"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3759  | 
by (drule pw_leq_step) (rule pw_leq_empty, simp)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3760  | 
|
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3761  | 
lemma pw_leq_split:  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3762  | 
assumes "pw_leq X Y"  | 
| 60495 | 3763  | 
  shows "\<exists>A B Z. X = A + Z \<and> Y = B + Z \<and> ((set_mset A, set_mset B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3764  | 
using assms  | 
| 60606 | 3765  | 
proof induct  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3766  | 
case pw_leq_empty thus ?case by auto  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3767  | 
next  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3768  | 
case (pw_leq_step x y X Y)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3769  | 
then obtain A B Z where  | 
| 58425 | 3770  | 
[simp]: "X = A + Z" "Y = B + Z"  | 
| 60495 | 3771  | 
      and 1[simp]: "(set_mset A, set_mset B) \<in> max_strict \<or> (B = {#} \<and> A = {#})"
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3772  | 
by auto  | 
| 60606 | 3773  | 
from pw_leq_step consider "x = y" | "(x, y) \<in> pair_less"  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3774  | 
unfolding pair_leq_def by auto  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3775  | 
thus ?case  | 
| 60606 | 3776  | 
proof cases  | 
3777  | 
case [simp]: 1  | 
|
3778  | 
    have "{#x#} + X = A + ({#y#}+Z) \<and> {#y#} + Y = B + ({#y#}+Z) \<and>
 | 
|
3779  | 
      ((set_mset A, set_mset B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
 | 
|
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
3780  | 
by auto  | 
| 60606 | 3781  | 
thus ?thesis by blast  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3782  | 
next  | 
| 60606 | 3783  | 
case 2  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3784  | 
    let ?A' = "{#x#} + A" and ?B' = "{#y#} + B"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3785  | 
    have "{#x#} + X = ?A' + Z"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3786  | 
      "{#y#} + Y = ?B' + Z"
 | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
3787  | 
by auto  | 
| 58425 | 3788  | 
moreover have  | 
| 60495 | 3789  | 
"(set_mset ?A', set_mset ?B') \<in> max_strict"  | 
| 60606 | 3790  | 
using 1 2 unfolding max_strict_def  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3791  | 
by (auto elim!: max_ext.cases)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3792  | 
ultimately show ?thesis by blast  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3793  | 
qed  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3794  | 
qed  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3795  | 
|
| 58425 | 3796  | 
lemma  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3797  | 
assumes pwleq: "pw_leq Z Z'"  | 
| 60495 | 3798  | 
shows ms_strictI: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_strict"  | 
| 60606 | 3799  | 
and ms_weakI1: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_weak"  | 
3800  | 
    and ms_weakI2:  "(Z + {#}, Z' + {#}) \<in> ms_weak"
 | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3801  | 
proof -  | 
| 58425 | 3802  | 
from pw_leq_split[OF pwleq]  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3803  | 
obtain A' B' Z''  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3804  | 
where [simp]: "Z = A' + Z''" "Z' = B' + Z''"  | 
| 60495 | 3805  | 
    and mx_or_empty: "(set_mset A', set_mset B') \<in> max_strict \<or> (A' = {#} \<and> B' = {#})"
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3806  | 
by blast  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3807  | 
  {
 | 
| 60495 | 3808  | 
assume max: "(set_mset A, set_mset B) \<in> max_strict"  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3809  | 
from mx_or_empty  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3810  | 
have "(Z'' + (A + A'), Z'' + (B + B')) \<in> ms_strict"  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3811  | 
proof  | 
| 60495 | 3812  | 
assume max': "(set_mset A', set_mset B') \<in> max_strict"  | 
3813  | 
with max have "(set_mset (A + A'), set_mset (B + B')) \<in> max_strict"  | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3814  | 
by (auto simp: max_strict_def intro: max_ext_additive)  | 
| 58425 | 3815  | 
thus ?thesis by (rule smsI)  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3816  | 
next  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3817  | 
      assume [simp]: "A' = {#} \<and> B' = {#}"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3818  | 
show ?thesis by (rule smsI) (auto intro: max)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3819  | 
qed  | 
| 60606 | 3820  | 
thus "(Z + A, Z' + B) \<in> ms_strict" by (simp add: ac_simps)  | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3821  | 
thus "(Z + A, Z' + B) \<in> ms_weak" by (simp add: ms_weak_def)  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3822  | 
}  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3823  | 
from mx_or_empty  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3824  | 
have "(Z'' + A', Z'' + B') \<in> ms_weak" by (rule wmsI)  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
3825  | 
  thus "(Z + {#}, Z' + {#}) \<in> ms_weak" by (simp add: ac_simps)
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3826  | 
qed  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3827  | 
|
| 39301 | 3828  | 
lemma empty_neutral: "{#} + x = x" "x + {#} = x"
 | 
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3829  | 
and nonempty_plus: "{# x #} + rs \<noteq> {#}"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3830  | 
and nonempty_single: "{# x #} \<noteq> {#}"
 | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3831  | 
by auto  | 
| 
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3832  | 
|
| 60500 | 3833  | 
setup \<open>  | 
| 60606 | 3834  | 
let  | 
| 74634 | 3835  | 
fun msetT T = \<^Type>\<open>multiset T\<close>;  | 
3836  | 
||
3837  | 
    fun mk_mset T [] = \<^instantiate>\<open>'a = T in term \<open>{#}\<close>\<close>
 | 
|
3838  | 
      | mk_mset T [x] = \<^instantiate>\<open>'a = T and x in term \<open>{#x#}\<close>\<close>
 | 
|
3839  | 
| mk_mset T (x :: xs) = \<^Const>\<open>plus \<open>msetT T\<close> for \<open>mk_mset T [x]\<close> \<open>mk_mset T xs\<close>\<close>  | 
|
| 60606 | 3840  | 
|
| 60752 | 3841  | 
fun mset_member_tac ctxt m i =  | 
| 60606 | 3842  | 
if m <= 0 then  | 
| 60752 | 3843  | 
        resolve_tac ctxt @{thms multi_member_this} i ORELSE
 | 
3844  | 
        resolve_tac ctxt @{thms multi_member_last} i
 | 
|
| 60606 | 3845  | 
else  | 
| 60752 | 3846  | 
        resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i
 | 
3847  | 
||
3848  | 
fun mset_nonempty_tac ctxt =  | 
|
3849  | 
      resolve_tac ctxt @{thms nonempty_plus} ORELSE'
 | 
|
3850  | 
      resolve_tac ctxt @{thms nonempty_single}
 | 
|
| 
29125
 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 
krauss 
parents: 
28823 
diff
changeset
 | 
3851  | 
|
| 60606 | 3852  | 
fun regroup_munion_conv ctxt =  | 
| 73393 | 3853  | 
Function_Lib.regroup_conv ctxt \<^const_abbrev>\<open>empty_mset\<close> \<^const_name>\<open>plus\<close>  | 
| 60606 | 3854  | 
        (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral}))
 | 
3855  | 
||
| 60752 | 3856  | 
fun unfold_pwleq_tac ctxt i =  | 
3857  | 
      (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st))
 | 
|
3858  | 
        ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i)
 | 
|
3859  | 
        ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i)
 | 
|
| 60606 | 3860  | 
|
3861  | 
    val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union},
 | 
|
3862  | 
                        @{thm Un_insert_left}, @{thm Un_empty_left}]
 | 
|
3863  | 
in  | 
|
3864  | 
ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset  | 
|
3865  | 
    {
 | 
|
3866  | 
msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv,  | 
|
3867  | 
mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac,  | 
|
3868  | 
mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps,  | 
|
3869  | 
      smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1},
 | 
|
| 60752 | 3870  | 
      reduction_pair = @{thm ms_reduction_pair}
 | 
| 60606 | 3871  | 
})  | 
3872  | 
end  | 
|
| 60500 | 3873  | 
\<close>  | 
3874  | 
||
3875  | 
||
3876  | 
subsection \<open>Legacy theorem bindings\<close>  | 
|
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3877  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39301 
diff
changeset
 | 
3878  | 
lemmas multi_count_eq = multiset_eq_iff [symmetric]  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3879  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3880  | 
lemma union_commute: "M + N = N + (M::'a multiset)"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
3881  | 
by (fact add.commute)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3882  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3883  | 
lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
3884  | 
by (fact add.assoc)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3885  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3886  | 
lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))"  | 
| 
57512
 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 
haftmann 
parents: 
57492 
diff
changeset
 | 
3887  | 
by (fact add.left_commute)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3888  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3889  | 
lemmas union_ac = union_assoc union_commute union_lcomm add_mset_commute  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3890  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3891  | 
lemma union_right_cancel: "M + K = N + K \<longleftrightarrow> M = (N::'a multiset)"  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3892  | 
by (fact add_right_cancel)  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3893  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3894  | 
lemma union_left_cancel: "K + M = K + N \<longleftrightarrow> M = (N::'a multiset)"  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3895  | 
by (fact add_left_cancel)  | 
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3896  | 
|
| 
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3897  | 
lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y"  | 
| 59557 | 3898  | 
by (fact add_left_imp_eq)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3899  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
3900  | 
lemma mset_subset_trans: "(M::'a multiset) \<subset># K \<Longrightarrow> K \<subset># N \<Longrightarrow> M \<subset># N"  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
3901  | 
by (fact subset_mset.less_trans)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3902  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3903  | 
lemma multiset_inter_commute: "A \<inter># B = B \<inter># A"  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
3904  | 
by (fact subset_mset.inf.commute)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3905  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3906  | 
lemma multiset_inter_assoc: "A \<inter># (B \<inter># C) = A \<inter># B \<inter># C"  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
3907  | 
by (fact subset_mset.inf.assoc [symmetric])  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3908  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3909  | 
lemma multiset_inter_left_commute: "A \<inter># (B \<inter># C) = B \<inter># (A \<inter># C)"  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
3910  | 
by (fact subset_mset.inf.left_commute)  | 
| 
35268
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3911  | 
|
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3912  | 
lemmas multiset_inter_ac =  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3913  | 
multiset_inter_commute  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3914  | 
multiset_inter_assoc  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3915  | 
multiset_inter_left_commute  | 
| 
 
04673275441a
switched notations for pointwise and multiset order
 
haftmann 
parents: 
35028 
diff
changeset
 | 
3916  | 
|
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3917  | 
lemma mset_le_not_refl: "\<not> M < (M::'a::preorder multiset)"  | 
| 
63388
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3918  | 
by (fact less_irrefl)  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3919  | 
|
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3920  | 
lemma mset_le_trans: "K < M \<Longrightarrow> M < N \<Longrightarrow> K < (N::'a::preorder multiset)"  | 
| 
63388
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3921  | 
by (fact less_trans)  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3922  | 
|
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3923  | 
lemma mset_le_not_sym: "M < N \<Longrightarrow> \<not> N < (M::'a::preorder multiset)"  | 
| 
63388
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3924  | 
by (fact less_not_sym)  | 
| 
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3925  | 
|
| 
63410
 
9789ccc2a477
more instantiations for multiset
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63409 
diff
changeset
 | 
3926  | 
lemma mset_le_asym: "M < N \<Longrightarrow> (\<not> P \<Longrightarrow> N < (M::'a::preorder multiset)) \<Longrightarrow> P"  | 
| 
63388
 
a095acd4cfbf
instantiate multiset with multiset ordering
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63360 
diff
changeset
 | 
3927  | 
by (fact less_asym)  | 
| 
34943
 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 
haftmann 
parents: 
33102 
diff
changeset
 | 
3928  | 
|
| 60500 | 3929  | 
declaration \<open>  | 
| 60606 | 3930  | 
let  | 
3931  | 
fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') =  | 
|
3932  | 
let  | 
|
3933  | 
val (maybe_opt, ps) =  | 
|
3934  | 
Nitpick_Model.dest_plain_fun t'  | 
|
| 67398 | 3935  | 
||> (~~)  | 
| 60606 | 3936  | 
||> map (apsnd (snd o HOLogic.dest_number))  | 
3937  | 
fun elems_for t =  | 
|
| 67398 | 3938  | 
(case AList.lookup (=) ps t of  | 
| 60606 | 3939  | 
SOME n => replicate n t  | 
3940  | 
| NONE => [Const (maybe_name, elem_T --> elem_T) $ t])  | 
|
3941  | 
in  | 
|
3942  | 
(case maps elems_for (all_values elem_T) @  | 
|
| 61333 | 3943  | 
(if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of  | 
| 74634 | 3944  | 
[] => \<^Const>\<open>Groups.zero T\<close>  | 
3945  | 
| ts => foldl1 (fn (s, t) => \<^Const>\<open>add_mset elem_T for s t\<close>) ts)  | 
|
| 60606 | 3946  | 
end  | 
3947  | 
| multiset_postproc _ _ _ _ t = t  | 
|
| 69593 | 3948  | 
in Nitpick_Model.register_term_postprocessor \<^typ>\<open>'a multiset\<close> multiset_postproc end  | 
| 60500 | 3949  | 
\<close>  | 
3950  | 
||
3951  | 
||
3952  | 
subsection \<open>Naive implementation using lists\<close>  | 
|
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3953  | 
|
| 60515 | 3954  | 
code_datatype mset  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3955  | 
|
| 60606 | 3956  | 
lemma [code]: "{#} = mset []"
 | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3957  | 
by simp  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3958  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3959  | 
lemma [code]: "add_mset x (mset xs) = mset (x # xs)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3960  | 
by simp  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3961  | 
|
| 63195 | 3962  | 
lemma [code]: "Multiset.is_empty (mset xs) \<longleftrightarrow> List.null xs"  | 
3963  | 
by (simp add: Multiset.is_empty_def List.null_def)  | 
|
3964  | 
||
| 60606 | 3965  | 
lemma union_code [code]: "mset xs + mset ys = mset (xs @ ys)"  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3966  | 
by simp  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3967  | 
|
| 60606 | 3968  | 
lemma [code]: "image_mset f (mset xs) = mset (map f xs)"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3969  | 
by simp  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3970  | 
|
| 60606 | 3971  | 
lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)"  | 
| 69442 | 3972  | 
by simp  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3973  | 
|
| 60606 | 3974  | 
lemma [code]: "mset xs - mset ys = mset (fold remove1 ys xs)"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
3975  | 
by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute diff_diff_add)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3976  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3977  | 
lemma [code]:  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3978  | 
"mset xs \<inter># mset ys =  | 
| 60515 | 3979  | 
mset (snd (fold (\<lambda>x (ys, zs).  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3980  | 
if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, [])))"  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3981  | 
proof -  | 
| 60515 | 3982  | 
have "\<And>zs. mset (snd (fold (\<lambda>x (ys, zs).  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3983  | 
if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) =  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3984  | 
(mset xs \<inter># mset ys) + mset zs"  | 
| 51623 | 3985  | 
by (induct xs arbitrary: ys)  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
3986  | 
(auto simp add: inter_add_right1 inter_add_right2 ac_simps)  | 
| 51623 | 3987  | 
then show ?thesis by simp  | 
3988  | 
qed  | 
|
3989  | 
||
3990  | 
lemma [code]:  | 
|
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3991  | 
"mset xs \<union># mset ys =  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61378 
diff
changeset
 | 
3992  | 
mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))"  | 
| 51623 | 3993  | 
proof -  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61378 
diff
changeset
 | 
3994  | 
have "\<And>zs. mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) =  | 
| 
63919
 
9aed2da07200
# after multiset intersection and union symbol
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63908 
diff
changeset
 | 
3995  | 
(mset xs \<union># mset ys) + mset zs"  | 
| 51623 | 3996  | 
by (induct xs arbitrary: ys) (simp_all add: multiset_eq_iff)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3997  | 
then show ?thesis by simp  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3998  | 
qed  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
3999  | 
|
| 59813 | 4000  | 
declare in_multiset_in_set [code_unfold]  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4001  | 
|
| 60606 | 4002  | 
lemma [code]: "count (mset xs) x = fold (\<lambda>y. if x = y then Suc else id) xs 0"  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4003  | 
proof -  | 
| 60515 | 4004  | 
have "\<And>n. fold (\<lambda>y. if x = y then Suc else id) xs n = count (mset xs) x + n"  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4005  | 
by (induct xs) simp_all  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4006  | 
then show ?thesis by simp  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4007  | 
qed  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4008  | 
|
| 60515 | 4009  | 
declare set_mset_mset [code]  | 
4010  | 
||
4011  | 
declare sorted_list_of_multiset_mset [code]  | 
|
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4012  | 
|
| 61585 | 4013  | 
lemma [code]: \<comment> \<open>not very efficient, but representation-ignorant!\<close>  | 
| 60515 | 4014  | 
"mset_set A = mset (sorted_list_of_set A)"  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4015  | 
apply (cases "finite A")  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4016  | 
apply simp_all  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4017  | 
apply (induct A rule: finite_induct)  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
4018  | 
apply simp_all  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4019  | 
done  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4020  | 
|
| 60515 | 4021  | 
declare size_mset [code]  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4022  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4023  | 
fun subset_eq_mset_impl :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool option" where  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4024  | 
"subset_eq_mset_impl [] ys = Some (ys \<noteq> [])"  | 
| 67398 | 4025  | 
| "subset_eq_mset_impl (Cons x xs) ys = (case List.extract ((=) x) ys of  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4026  | 
None \<Rightarrow> None  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4027  | 
| Some (ys1,_,ys2) \<Rightarrow> subset_eq_mset_impl xs (ys1 @ ys2))"  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4028  | 
|
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4029  | 
lemma subset_eq_mset_impl: "(subset_eq_mset_impl xs ys = None \<longleftrightarrow> \<not> mset xs \<subseteq># mset ys) \<and>  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4030  | 
(subset_eq_mset_impl xs ys = Some True \<longleftrightarrow> mset xs \<subset># mset ys) \<and>  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4031  | 
(subset_eq_mset_impl xs ys = Some False \<longrightarrow> mset xs = mset ys)"  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4032  | 
proof (induct xs arbitrary: ys)  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4033  | 
case (Nil ys)  | 
| 64076 | 4034  | 
show ?case by (auto simp: subset_mset.zero_less_iff_neq_zero)  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4035  | 
next  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4036  | 
case (Cons x xs ys)  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4037  | 
show ?case  | 
| 67398 | 4038  | 
proof (cases "List.extract ((=) x) ys")  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4039  | 
case None  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4040  | 
hence x: "x \<notin> set ys" by (simp add: extract_None_iff)  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4041  | 
    {
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
4042  | 
assume "mset (x # xs) \<subseteq># mset ys"  | 
| 60495 | 4043  | 
from set_mset_mono[OF this] x have False by simp  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4044  | 
} note nle = this  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4045  | 
moreover  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4046  | 
    {
 | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
4047  | 
assume "mset (x # xs) \<subset># mset ys"  | 
| 
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
4048  | 
hence "mset (x # xs) \<subseteq># mset ys" by auto  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4049  | 
from nle[OF this] have False .  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4050  | 
}  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4051  | 
ultimately show ?thesis using None by auto  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4052  | 
next  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4053  | 
case (Some res)  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4054  | 
obtain ys1 y ys2 where res: "res = (ys1,y,ys2)" by (cases res, auto)  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4055  | 
note Some = Some[unfolded res]  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4056  | 
from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4057  | 
hence id: "mset ys = add_mset x (mset (ys1 @ ys2))"  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
4058  | 
by auto  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4059  | 
show ?thesis unfolding subset_eq_mset_impl.simps  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4060  | 
unfolding Some option.simps split  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4061  | 
unfolding id  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4062  | 
using Cons[of "ys1 @ ys2"]  | 
| 
60397
 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
59999 
diff
changeset
 | 
4063  | 
unfolding subset_mset_def subseteq_mset_def by auto  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4064  | 
qed  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4065  | 
qed  | 
| 
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4066  | 
|
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4067  | 
lemma [code]: "mset xs \<subseteq># mset ys \<longleftrightarrow> subset_eq_mset_impl xs ys \<noteq> None"  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4068  | 
using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto)  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4069  | 
|
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4070  | 
lemma [code]: "mset xs \<subset># mset ys \<longleftrightarrow> subset_eq_mset_impl xs ys = Some True"  | 
| 
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4071  | 
using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4072  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4073  | 
instantiation multiset :: (equal) equal  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4074  | 
begin  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4075  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4076  | 
definition  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4077  | 
[code del]: "HOL.equal A (B :: 'a multiset) \<longleftrightarrow> A = B"  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4078  | 
lemma [code]: "HOL.equal (mset xs) (mset ys) \<longleftrightarrow> subset_eq_mset_impl xs ys = Some False"  | 
| 
55808
 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 
nipkow 
parents: 
55565 
diff
changeset
 | 
4079  | 
unfolding equal_multiset_def  | 
| 
63310
 
caaacf37943f
normalising multiset theorem names
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63290 
diff
changeset
 | 
4080  | 
using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4081  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4082  | 
instance  | 
| 60678 | 4083  | 
by standard (simp add: equal_multiset_def)  | 
4084  | 
||
| 
37169
 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 
blanchet 
parents: 
37107 
diff
changeset
 | 
4085  | 
end  | 
| 49388 | 4086  | 
|
| 66313 | 4087  | 
declare sum_mset_sum_list [code]  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4088  | 
|
| 63830 | 4089  | 
lemma [code]: "prod_mset (mset xs) = fold times xs 1"  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4090  | 
proof -  | 
| 63830 | 4091  | 
have "\<And>x. fold times xs x = prod_mset (mset xs) * x"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4092  | 
by (induct xs) (simp_all add: ac_simps)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4093  | 
then show ?thesis by simp  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4094  | 
qed  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4095  | 
|
| 60500 | 4096  | 
text \<open>  | 
| 69593 | 4097  | 
Exercise for the casual reader: add implementations for \<^term>\<open>(\<le>)\<close>  | 
4098  | 
and \<^term>\<open>(<)\<close> (multiset order).  | 
|
| 60500 | 4099  | 
\<close>  | 
4100  | 
||
4101  | 
text \<open>Quickcheck generators\<close>  | 
|
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4102  | 
|
| 72607 | 4103  | 
context  | 
4104  | 
includes term_syntax  | 
|
4105  | 
begin  | 
|
4106  | 
||
4107  | 
definition  | 
|
| 61076 | 4108  | 
msetify :: "'a::typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term)  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4109  | 
\<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where  | 
| 60515 | 4110  | 
  [code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\<cdot>} xs"
 | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4111  | 
|
| 72607 | 4112  | 
end  | 
4113  | 
||
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4114  | 
instantiation multiset :: (random) random  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4115  | 
begin  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4116  | 
|
| 72581 | 4117  | 
context  | 
4118  | 
includes state_combinator_syntax  | 
|
4119  | 
begin  | 
|
4120  | 
||
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4121  | 
definition  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4122  | 
"Quickcheck_Random.random i = Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (msetify xs))"  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4123  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4124  | 
instance ..  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4125  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4126  | 
end  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4127  | 
|
| 72581 | 4128  | 
end  | 
| 
51600
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4129  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4130  | 
instantiation multiset :: (full_exhaustive) full_exhaustive  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4131  | 
begin  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4132  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4133  | 
definition full_exhaustive_multiset :: "('a multiset \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option"
 | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4134  | 
where  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4135  | 
"full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (\<lambda>xs. f (msetify xs)) i"  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4136  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4137  | 
instance ..  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4138  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4139  | 
end  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4140  | 
|
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4141  | 
hide_const (open) msetify  | 
| 
 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 
haftmann 
parents: 
51599 
diff
changeset
 | 
4142  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4143  | 
|
| 60500 | 4144  | 
subsection \<open>BNF setup\<close>  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4145  | 
|
| 57966 | 4146  | 
definition rel_mset where  | 
| 60515 | 4147  | 
"rel_mset R X Y \<longleftrightarrow> (\<exists>xs ys. mset xs = X \<and> mset ys = Y \<and> list_all2 R xs ys)"  | 
4148  | 
||
4149  | 
lemma mset_zip_take_Cons_drop_twice:  | 
|
| 57966 | 4150  | 
assumes "length xs = length ys" "j \<le> length xs"  | 
| 60515 | 4151  | 
shows "mset (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) =  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4152  | 
add_mset (x,y) (mset (zip xs ys))"  | 
| 60606 | 4153  | 
using assms  | 
| 57966 | 4154  | 
proof (induct xs ys arbitrary: x y j rule: list_induct2)  | 
4155  | 
case Nil  | 
|
4156  | 
thus ?case  | 
|
4157  | 
by simp  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4158  | 
next  | 
| 57966 | 4159  | 
case (Cons x xs y ys)  | 
4160  | 
thus ?case  | 
|
4161  | 
proof (cases "j = 0")  | 
|
4162  | 
case True  | 
|
4163  | 
thus ?thesis  | 
|
4164  | 
by simp  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4165  | 
next  | 
| 57966 | 4166  | 
case False  | 
4167  | 
then obtain k where k: "j = Suc k"  | 
|
| 60678 | 4168  | 
by (cases j) simp  | 
| 57966 | 4169  | 
hence "k \<le> length xs"  | 
4170  | 
using Cons.prems by auto  | 
|
| 60515 | 4171  | 
hence "mset (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) =  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4172  | 
add_mset (x,y) (mset (zip xs ys))"  | 
| 57966 | 4173  | 
by (rule Cons.hyps(2))  | 
4174  | 
thus ?thesis  | 
|
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
4175  | 
unfolding k by auto  | 
| 58425 | 4176  | 
qed  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4177  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4178  | 
|
| 60515 | 4179  | 
lemma ex_mset_zip_left:  | 
4180  | 
assumes "length xs = length ys" "mset xs' = mset xs"  | 
|
4181  | 
shows "\<exists>ys'. length ys' = length xs' \<and> mset (zip xs' ys') = mset (zip xs ys)"  | 
|
| 58425 | 4182  | 
using assms  | 
| 57966 | 4183  | 
proof (induct xs ys arbitrary: xs' rule: list_induct2)  | 
4184  | 
case Nil  | 
|
4185  | 
thus ?case  | 
|
4186  | 
by auto  | 
|
4187  | 
next  | 
|
4188  | 
case (Cons x xs y ys xs')  | 
|
4189  | 
obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x"  | 
|
| 60515 | 4190  | 
by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD)  | 
| 58425 | 4191  | 
|
| 63040 | 4192  | 
define xsa where "xsa = take j xs' @ drop (Suc j) xs'"  | 
| 60515 | 4193  | 
  have "mset xs' = {#x#} + mset xsa"
 | 
| 57966 | 4194  | 
unfolding xsa_def using j_len nth_j  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4195  | 
by (metis Cons_nth_drop_Suc union_mset_add_mset_right add_mset_remove_trivial add_diff_cancel_left'  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4196  | 
append_take_drop_id mset.simps(2) mset_append)  | 
| 60515 | 4197  | 
hence ms_x: "mset xsa = mset xs"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4198  | 
by (simp add: Cons.prems)  | 
| 57966 | 4199  | 
then obtain ysa where  | 
| 60515 | 4200  | 
len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)"  | 
| 57966 | 4201  | 
using Cons.hyps(2) by blast  | 
4202  | 
||
| 63040 | 4203  | 
define ys' where "ys' = take j ysa @ y # drop j ysa"  | 
| 57966 | 4204  | 
have xs': "xs' = take j xsa @ x # drop j xsa"  | 
4205  | 
using ms_x j_len nth_j Cons.prems xsa_def  | 
|
| 
58247
 
98d0f85d247f
enamed drop_Suc_conv_tl and nth_drop' to Cons_nth_drop_Suc
 
nipkow 
parents: 
58098 
diff
changeset
 | 
4206  | 
by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons  | 
| 60515 | 4207  | 
length_drop size_mset)  | 
| 57966 | 4208  | 
have j_len': "j \<le> length xsa"  | 
4209  | 
using j_len xs' xsa_def  | 
|
4210  | 
by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less)  | 
|
4211  | 
have "length ys' = length xs'"  | 
|
4212  | 
unfolding ys'_def using Cons.prems len_a ms_x  | 
|
| 60515 | 4213  | 
by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length)  | 
4214  | 
moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))"  | 
|
| 57966 | 4215  | 
unfolding xs' ys'_def  | 
| 60515 | 4216  | 
by (rule trans[OF mset_zip_take_Cons_drop_twice])  | 
| 
63794
 
bcec0534aeea
clean argument of simp add
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63793 
diff
changeset
 | 
4217  | 
(auto simp: len_a ms_a j_len')  | 
| 57966 | 4218  | 
ultimately show ?case  | 
4219  | 
by blast  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4220  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4221  | 
|
| 57966 | 4222  | 
lemma list_all2_reorder_left_invariance:  | 
| 60515 | 4223  | 
assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs"  | 
4224  | 
shows "\<exists>ys'. list_all2 R xs' ys' \<and> mset ys' = mset ys"  | 
|
| 57966 | 4225  | 
proof -  | 
4226  | 
have len: "length xs = length ys"  | 
|
4227  | 
using rel list_all2_conv_all_nth by auto  | 
|
4228  | 
obtain ys' where  | 
|
| 60515 | 4229  | 
len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)"  | 
4230  | 
using len ms_x by (metis ex_mset_zip_left)  | 
|
| 57966 | 4231  | 
have "list_all2 R xs' ys'"  | 
| 60515 | 4232  | 
using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD)  | 
4233  | 
moreover have "mset ys' = mset ys"  | 
|
4234  | 
using len len' ms_xy map_snd_zip mset_map by metis  | 
|
| 57966 | 4235  | 
ultimately show ?thesis  | 
4236  | 
by blast  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4237  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4238  | 
|
| 60515 | 4239  | 
lemma ex_mset: "\<exists>xs. mset xs = X"  | 
4240  | 
by (induct X) (simp, metis mset.simps(2))  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4241  | 
|
| 73301 | 4242  | 
inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool" 
 | 
| 62324 | 4243  | 
where  | 
4244  | 
  "pred_mset P {#}"
 | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4245  | 
| "\<lbrakk>P a; pred_mset P M\<rbrakk> \<Longrightarrow> pred_mset P (add_mset a M)"  | 
| 62324 | 4246  | 
|
| 73301 | 4247  | 
lemma pred_mset_iff: \<comment> \<open>TODO: alias for \<^const>\<open>Multiset.Ball\<close>\<close>  | 
4248  | 
\<open>pred_mset P M \<longleftrightarrow> Multiset.Ball M P\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>)  | 
|
4249  | 
proof  | 
|
4250  | 
assume ?P  | 
|
4251  | 
then show ?Q by induction simp_all  | 
|
4252  | 
next  | 
|
4253  | 
assume ?Q  | 
|
4254  | 
then show ?P  | 
|
4255  | 
by (induction M) (auto intro: pred_mset.intros)  | 
|
4256  | 
qed  | 
|
4257  | 
||
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4258  | 
bnf "'a multiset"  | 
| 57966 | 4259  | 
map: image_mset  | 
| 60495 | 4260  | 
sets: set_mset  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4261  | 
bd: natLeq  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4262  | 
  wits: "{#}"
 | 
| 57966 | 4263  | 
rel: rel_mset  | 
| 62324 | 4264  | 
pred: pred_mset  | 
| 57966 | 4265  | 
proof -  | 
4266  | 
show "image_mset id = id"  | 
|
4267  | 
by (rule image_mset.id)  | 
|
| 60606 | 4268  | 
show "image_mset (g \<circ> f) = image_mset g \<circ> image_mset f" for f g  | 
| 59813 | 4269  | 
unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality)  | 
| 60606 | 4270  | 
show "(\<And>z. z \<in> set_mset X \<Longrightarrow> f z = g z) \<Longrightarrow> image_mset f X = image_mset g X" for f g X  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
4271  | 
by (induct X) simp_all  | 
| 67398 | 4272  | 
show "set_mset \<circ> image_mset f = (`) f \<circ> set_mset" for f  | 
| 57966 | 4273  | 
by auto  | 
4274  | 
show "card_order natLeq"  | 
|
4275  | 
by (rule natLeq_card_order)  | 
|
4276  | 
show "BNF_Cardinal_Arithmetic.cinfinite natLeq"  | 
|
4277  | 
by (rule natLeq_cinfinite)  | 
|
| 75624 | 4278  | 
show "regularCard natLeq"  | 
4279  | 
by (rule regularCard_natLeq)  | 
|
4280  | 
show "ordLess2 (card_of (set_mset X)) natLeq" for X  | 
|
| 57966 | 4281  | 
by transfer  | 
| 75624 | 4282  | 
(auto simp: finite_iff_ordLess_natLeq[symmetric])  | 
| 60606 | 4283  | 
show "rel_mset R OO rel_mset S \<le> rel_mset (R OO S)" for R S  | 
| 57966 | 4284  | 
unfolding rel_mset_def[abs_def] OO_def  | 
4285  | 
apply clarify  | 
|
| 60678 | 4286  | 
subgoal for X Z Y xs ys' ys zs  | 
4287  | 
apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys])  | 
|
4288  | 
apply (auto intro: list_all2_trans)  | 
|
4289  | 
done  | 
|
| 60606 | 4290  | 
done  | 
4291  | 
show "rel_mset R =  | 
|
| 62324 | 4292  | 
    (\<lambda>x y. \<exists>z. set_mset z \<subseteq> {(x, y). R x y} \<and>
 | 
4293  | 
image_mset fst z = x \<and> image_mset snd z = y)" for R  | 
|
4294  | 
unfolding rel_mset_def[abs_def]  | 
|
| 57966 | 4295  | 
apply (rule ext)+  | 
| 62324 | 4296  | 
apply safe  | 
4297  | 
apply (rule_tac x = "mset (zip xs ys)" in exI;  | 
|
| 68406 | 4298  | 
auto simp: in_set_zip list_all2_iff simp flip: mset_map)  | 
| 57966 | 4299  | 
apply (rename_tac XY)  | 
| 60515 | 4300  | 
apply (cut_tac X = XY in ex_mset)  | 
| 57966 | 4301  | 
apply (erule exE)  | 
4302  | 
apply (rename_tac xys)  | 
|
4303  | 
apply (rule_tac x = "map fst xys" in exI)  | 
|
| 60515 | 4304  | 
apply (auto simp: mset_map)  | 
| 57966 | 4305  | 
apply (rule_tac x = "map snd xys" in exI)  | 
| 60515 | 4306  | 
apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd)  | 
| 59997 | 4307  | 
done  | 
| 60606 | 4308  | 
  show "z \<in> set_mset {#} \<Longrightarrow> False" for z
 | 
| 57966 | 4309  | 
by auto  | 
| 62324 | 4310  | 
show "pred_mset P = (\<lambda>x. Ball (set_mset x) P)" for P  | 
| 73301 | 4311  | 
by (simp add: fun_eq_iff pred_mset_iff)  | 
| 57966 | 4312  | 
qed  | 
4313  | 
||
| 73301 | 4314  | 
inductive rel_mset' :: \<open>('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset \<Rightarrow> bool\<close>
 | 
| 60606 | 4315  | 
where  | 
| 57966 | 4316  | 
  Zero[intro]: "rel_mset' R {#} {#}"
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4317  | 
| Plus[intro]: "\<lbrakk>R a b; rel_mset' R M N\<rbrakk> \<Longrightarrow> rel_mset' R (add_mset a M) (add_mset b N)"  | 
| 57966 | 4318  | 
|
4319  | 
lemma rel_mset_Zero: "rel_mset R {#} {#}"
 | 
|
4320  | 
unfolding rel_mset_def Grp_def by auto  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4321  | 
|
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4322  | 
declare multiset.count[simp]  | 
| 73270 | 4323  | 
declare count_Abs_multiset[simp]  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4324  | 
declare multiset.count_inverse[simp]  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4325  | 
|
| 57966 | 4326  | 
lemma rel_mset_Plus:  | 
| 60606 | 4327  | 
assumes ab: "R a b"  | 
4328  | 
and MN: "rel_mset R M N"  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4329  | 
shows "rel_mset R (add_mset a M) (add_mset b N)"  | 
| 60606 | 4330  | 
proof -  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4331  | 
have "\<exists>ya. add_mset a (image_mset fst y) = image_mset fst ya \<and>  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4332  | 
add_mset b (image_mset snd y) = image_mset snd ya \<and>  | 
| 60606 | 4333  | 
    set_mset ya \<subseteq> {(x, y). R x y}"
 | 
4334  | 
    if "R a b" and "set_mset y \<subseteq> {(x, y). R x y}" for y
 | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4335  | 
using that by (intro exI[of _ "add_mset (a,b) y"]) auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4336  | 
thus ?thesis  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4337  | 
using assms  | 
| 57966 | 4338  | 
unfolding multiset.rel_compp_Grp Grp_def by blast  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4339  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4340  | 
|
| 60606 | 4341  | 
lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \<Longrightarrow> rel_mset R M N"  | 
| 60678 | 4342  | 
by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus)  | 
| 57966 | 4343  | 
|
| 60606 | 4344  | 
lemma rel_mset_size: "rel_mset R M N \<Longrightarrow> size M = size N"  | 
| 60678 | 4345  | 
unfolding multiset.rel_compp_Grp Grp_def by auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4346  | 
|
| 73594 | 4347  | 
lemma rel_mset_Zero_iff [simp]:  | 
4348  | 
  shows "rel_mset rel {#} Y \<longleftrightarrow> Y = {#}" and "rel_mset rel X {#} \<longleftrightarrow> X = {#}"
 | 
|
4349  | 
by (auto simp add: rel_mset_Zero dest: rel_mset_size)  | 
|
4350  | 
||
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4351  | 
lemma multiset_induct2[case_names empty addL addR]:  | 
| 60678 | 4352  | 
  assumes empty: "P {#} {#}"
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4353  | 
and addL: "\<And>a M N. P M N \<Longrightarrow> P (add_mset a M) N"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4354  | 
and addR: "\<And>a M N. P M N \<Longrightarrow> P M (add_mset a N)"  | 
| 60678 | 4355  | 
shows "P M N"  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4356  | 
apply(induct N rule: multiset_induct)  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4357  | 
apply(induct M rule: multiset_induct, rule empty, erule addL)  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4358  | 
apply(induct M rule: multiset_induct, erule addR, erule addR)  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4359  | 
done  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4360  | 
|
| 59949 | 4361  | 
lemma multiset_induct2_size[consumes 1, case_names empty add]:  | 
| 60606 | 4362  | 
assumes c: "size M = size N"  | 
4363  | 
    and empty: "P {#} {#}"
 | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4364  | 
and add: "\<And>a b M N a b. P M N \<Longrightarrow> P (add_mset a M) (add_mset b N)"  | 
| 60606 | 4365  | 
shows "P M N"  | 
| 60678 | 4366  | 
using c  | 
4367  | 
proof (induct M arbitrary: N rule: measure_induct_rule[of size])  | 
|
| 60606 | 4368  | 
case (less M)  | 
4369  | 
show ?case  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4370  | 
  proof(cases "M = {#}")
 | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4371  | 
    case True hence "N = {#}" using less.prems by auto
 | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4372  | 
thus ?thesis using True empty by auto  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4373  | 
next  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4374  | 
case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split)  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4375  | 
    have "N \<noteq> {#}" using False less.prems by auto
 | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4376  | 
then obtain N1 b where N: "N = add_mset b N1" by (metis multi_nonempty_split)  | 
| 59949 | 4377  | 
have "size M1 = size N1" using less.prems unfolding M N by auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4378  | 
thus ?thesis using M N less.hyps add by auto  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4379  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4380  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4381  | 
|
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4382  | 
lemma msed_map_invL:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4383  | 
assumes "image_mset f (add_mset a M) = N"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4384  | 
shows "\<exists>N1. N = add_mset (f a) N1 \<and> image_mset f M = N1"  | 
| 60606 | 4385  | 
proof -  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4386  | 
have "f a \<in># N"  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4387  | 
using assms multiset.set_map[of f "add_mset a M"] by auto  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4388  | 
then obtain N1 where N: "N = add_mset (f a) N1" using multi_member_split by metis  | 
| 57966 | 4389  | 
have "image_mset f M = N1" using assms unfolding N by simp  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4390  | 
thus ?thesis using N by blast  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4391  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4392  | 
|
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4393  | 
lemma msed_map_invR:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4394  | 
assumes "image_mset f M = add_mset b N"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4395  | 
shows "\<exists>M1 a. M = add_mset a M1 \<and> f a = b \<and> image_mset f M1 = N"  | 
| 60606 | 4396  | 
proof -  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4397  | 
obtain a where a: "a \<in># M" and fa: "f a = b"  | 
| 60606 | 4398  | 
using multiset.set_map[of f M] unfolding assms  | 
| 
62430
 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 
haftmann 
parents: 
62390 
diff
changeset
 | 
4399  | 
by (metis image_iff union_single_eq_member)  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4400  | 
then obtain M1 where M: "M = add_mset a M1" using multi_member_split by metis  | 
| 57966 | 4401  | 
have "image_mset f M1 = N" using assms unfolding M fa[symmetric] by simp  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4402  | 
thus ?thesis using M fa by blast  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4403  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4404  | 
|
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4405  | 
lemma msed_rel_invL:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4406  | 
assumes "rel_mset R (add_mset a M) N"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4407  | 
shows "\<exists>N1 b. N = add_mset b N1 \<and> R a b \<and> rel_mset R M N1"  | 
| 60606 | 4408  | 
proof -  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4409  | 
obtain K where KM: "image_mset fst K = add_mset a M"  | 
| 60606 | 4410  | 
    and KN: "image_mset snd K = N" and sK: "set_mset K \<subseteq> {(a, b). R a b}"
 | 
4411  | 
using assms  | 
|
4412  | 
unfolding multiset.rel_compp_Grp Grp_def by auto  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4413  | 
obtain K1 ab where K: "K = add_mset ab K1" and a: "fst ab = a"  | 
| 60606 | 4414  | 
and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4415  | 
obtain N1 where N: "N = add_mset (snd ab) N1" and K1N1: "image_mset snd K1 = N1"  | 
| 60606 | 4416  | 
using msed_map_invL[OF KN[unfolded K]] by auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4417  | 
have Rab: "R a (snd ab)" using sK a unfolding K by auto  | 
| 57966 | 4418  | 
have "rel_mset R M N1" using sK K1M K1N1  | 
| 60606 | 4419  | 
unfolding K multiset.rel_compp_Grp Grp_def by auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4420  | 
thus ?thesis using N Rab by auto  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4421  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4422  | 
|
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4423  | 
lemma msed_rel_invR:  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4424  | 
assumes "rel_mset R M (add_mset b N)"  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4425  | 
shows "\<exists>M1 a. M = add_mset a M1 \<and> R a b \<and> rel_mset R M1 N"  | 
| 60606 | 4426  | 
proof -  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4427  | 
obtain K where KN: "image_mset snd K = add_mset b N"  | 
| 60606 | 4428  | 
    and KM: "image_mset fst K = M" and sK: "set_mset K \<subseteq> {(a, b). R a b}"
 | 
4429  | 
using assms  | 
|
4430  | 
unfolding multiset.rel_compp_Grp Grp_def by auto  | 
|
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4431  | 
obtain K1 ab where K: "K = add_mset ab K1" and b: "snd ab = b"  | 
| 60606 | 4432  | 
and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4433  | 
obtain M1 where M: "M = add_mset (fst ab) M1" and K1M1: "image_mset fst K1 = M1"  | 
| 60606 | 4434  | 
using msed_map_invL[OF KM[unfolded K]] by auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4435  | 
have Rab: "R (fst ab) b" using sK b unfolding K by auto  | 
| 57966 | 4436  | 
have "rel_mset R M1 N" using sK K1N K1M1  | 
| 60606 | 4437  | 
unfolding K multiset.rel_compp_Grp Grp_def by auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4438  | 
thus ?thesis using M Rab by auto  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4439  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4440  | 
|
| 57966 | 4441  | 
lemma rel_mset_imp_rel_mset':  | 
| 60606 | 4442  | 
assumes "rel_mset R M N"  | 
4443  | 
shows "rel_mset' R M N"  | 
|
| 59949 | 4444  | 
using assms proof(induct M arbitrary: N rule: measure_induct_rule[of size])  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4445  | 
case (less M)  | 
| 59949 | 4446  | 
have c: "size M = size N" using rel_mset_size[OF less.prems] .  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4447  | 
show ?case  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4448  | 
  proof(cases "M = {#}")
 | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4449  | 
    case True hence "N = {#}" using c by simp
 | 
| 57966 | 4450  | 
thus ?thesis using True rel_mset'.Zero by auto  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4451  | 
next  | 
| 
63793
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4452  | 
case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split)  | 
| 
 
e68a0b651eb5
add_mset constructor in multisets
 
fleury <Mathias.Fleury@mpi-inf.mpg.de> 
parents: 
63689 
diff
changeset
 | 
4453  | 
obtain N1 b where N: "N = add_mset b N1" and R: "R a b" and ms: "rel_mset R M1 N1"  | 
| 60606 | 4454  | 
using msed_rel_invL[OF less.prems[unfolded M]] by auto  | 
| 57966 | 4455  | 
have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp  | 
4456  | 
thus ?thesis using rel_mset'.Plus[of R a b, OF R] unfolding M N by simp  | 
|
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4457  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4458  | 
qed  | 
| 
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4459  | 
|
| 60606 | 4460  | 
lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N"  | 
| 60678 | 4461  | 
using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto  | 
| 57966 | 4462  | 
|
| 69593 | 4463  | 
text \<open>The main end product for \<^const>\<open>rel_mset\<close>: inductive characterization:\<close>  | 
| 61337 | 4464  | 
lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] =  | 
| 60606 | 4465  | 
rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]]  | 
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4466  | 
|
| 56656 | 4467  | 
|
| 60500 | 4468  | 
subsection \<open>Size setup\<close>  | 
| 56656 | 4469  | 
|
| 67332 | 4470  | 
lemma size_multiset_o_map: "size_multiset g \<circ> image_mset f = size_multiset (g \<circ> f)"  | 
| 65547 | 4471  | 
apply (rule ext)  | 
4472  | 
subgoal for x by (induct x) auto  | 
|
4473  | 
done  | 
|
| 56656 | 4474  | 
|
| 60500 | 4475  | 
setup \<open>  | 
| 69593 | 4476  | 
BNF_LFP_Size.register_size_global \<^type_name>\<open>multiset\<close> \<^const_name>\<open>size_multiset\<close>  | 
| 62082 | 4477  | 
    @{thm size_multiset_overloaded_def}
 | 
| 60606 | 4478  | 
    @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single
 | 
4479  | 
size_union}  | 
|
| 67332 | 4480  | 
    @{thms size_multiset_o_map}
 | 
| 60500 | 4481  | 
\<close>  | 
| 56656 | 4482  | 
|
| 65547 | 4483  | 
subsection \<open>Lemmas about Size\<close>  | 
4484  | 
||
4485  | 
lemma size_mset_SucE: "size A = Suc n \<Longrightarrow> (\<And>a B. A = {#a#} + B \<Longrightarrow> size B = n \<Longrightarrow> P) \<Longrightarrow> P"
 | 
|
4486  | 
by (cases A) (auto simp add: ac_simps)  | 
|
4487  | 
||
4488  | 
lemma size_Suc_Diff1: "x \<in># M \<Longrightarrow> Suc (size (M - {#x#})) = size M"
 | 
|
4489  | 
using arg_cong[OF insert_DiffM, of _ _ size] by simp  | 
|
4490  | 
||
4491  | 
lemma size_Diff_singleton: "x \<in># M \<Longrightarrow> size (M - {#x#}) = size M - 1"
 | 
|
| 68406 | 4492  | 
by (simp flip: size_Suc_Diff1)  | 
| 65547 | 4493  | 
|
4494  | 
lemma size_Diff_singleton_if: "size (A - {#x#}) = (if x \<in># A then size A - 1 else size A)"
 | 
|
4495  | 
by (simp add: diff_single_trivial size_Diff_singleton)  | 
|
4496  | 
||
4497  | 
lemma size_Un_Int: "size A + size B = size (A \<union># B) + size (A \<inter># B)"  | 
|
4498  | 
by (metis inter_subset_eq_union size_union subset_mset.diff_add union_diff_inter_eq_sup)  | 
|
4499  | 
||
4500  | 
lemma size_Un_disjoint: "A \<inter># B = {#} \<Longrightarrow> size (A \<union># B) = size A + size B"
 | 
|
4501  | 
using size_Un_Int[of A B] by simp  | 
|
4502  | 
||
4503  | 
lemma size_Diff_subset_Int: "size (M - M') = size M - size (M \<inter># M')"  | 
|
4504  | 
by (metis diff_intersect_left_idem size_Diff_submset subset_mset.inf_le1)  | 
|
4505  | 
||
4506  | 
lemma diff_size_le_size_Diff: "size (M :: _ multiset) - size M' \<le> size (M - M')"  | 
|
4507  | 
by (simp add: diff_le_mono2 size_Diff_subset_Int size_mset_mono)  | 
|
4508  | 
||
4509  | 
lemma size_Diff1_less: "x\<in># M \<Longrightarrow> size (M - {#x#}) < size M"
 | 
|
4510  | 
by (rule Suc_less_SucD) (simp add: size_Suc_Diff1)  | 
|
4511  | 
||
4512  | 
lemma size_Diff2_less: "x\<in># M \<Longrightarrow> y\<in># M \<Longrightarrow> size (M - {#x#} - {#y#}) < size M"
 | 
|
4513  | 
by (metis less_imp_diff_less size_Diff1_less size_Diff_subset_Int)  | 
|
4514  | 
||
4515  | 
lemma size_Diff1_le: "size (M - {#x#}) \<le> size M"
 | 
|
4516  | 
by (cases "x \<in># M") (simp_all add: size_Diff1_less less_imp_le diff_single_trivial)  | 
|
4517  | 
||
4518  | 
lemma size_psubset: "M \<subseteq># M' \<Longrightarrow> size M < size M' \<Longrightarrow> M \<subset># M'"  | 
|
4519  | 
using less_irrefl subset_mset_def by blast  | 
|
4520  | 
||
| 
76700
 
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
 
blanchet 
parents: 
76682 
diff
changeset
 | 
4521  | 
lifting_update multiset.lifting  | 
| 
 
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
 
blanchet 
parents: 
76682 
diff
changeset
 | 
4522  | 
lifting_forget multiset.lifting  | 
| 
 
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
 
blanchet 
parents: 
76682 
diff
changeset
 | 
4523  | 
|
| 56656 | 4524  | 
hide_const (open) wcount  | 
4525  | 
||
| 
55129
 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 
blanchet 
parents: 
54868 
diff
changeset
 | 
4526  | 
end  |