author | wenzelm |
Mon, 11 Apr 2016 15:26:58 +0200 | |
changeset 62954 | c5d0fdc260fa |
parent 62430 | 9527ff088c15 |
child 63040 | eb4ddd18d635 |
permissions | -rw-r--r-- |
59813 | 1 |
(* Title: HOL/Library/Multiset_Order.thy |
2 |
Author: Dmitriy Traytel, TU Muenchen |
|
3 |
Author: Jasmin Blanchette, Inria, LORIA, MPII |
|
4 |
*) |
|
5 |
||
60500 | 6 |
section \<open>More Theorems about the Multiset Order\<close> |
59813 | 7 |
|
8 |
theory Multiset_Order |
|
9 |
imports Multiset |
|
10 |
begin |
|
11 |
||
60500 | 12 |
subsubsection \<open>Alternative characterizations\<close> |
59813 | 13 |
|
14 |
context order |
|
15 |
begin |
|
16 |
||
17 |
lemma reflp_le: "reflp (op \<le>)" |
|
18 |
unfolding reflp_def by simp |
|
19 |
||
20 |
lemma antisymP_le: "antisymP (op \<le>)" |
|
21 |
unfolding antisym_def by auto |
|
22 |
||
23 |
lemma transp_le: "transp (op \<le>)" |
|
24 |
unfolding transp_def by auto |
|
25 |
||
26 |
lemma irreflp_less: "irreflp (op <)" |
|
27 |
unfolding irreflp_def by simp |
|
28 |
||
29 |
lemma antisymP_less: "antisymP (op <)" |
|
30 |
unfolding antisym_def by auto |
|
31 |
||
32 |
lemma transp_less: "transp (op <)" |
|
33 |
unfolding transp_def by auto |
|
34 |
||
35 |
lemmas le_trans = transp_le[unfolded transp_def, rule_format] |
|
36 |
||
37 |
lemma order_mult: "class.order |
|
38 |
(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y} \<or> M = N) |
|
39 |
(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y})" |
|
40 |
(is "class.order ?le ?less") |
|
41 |
proof - |
|
42 |
have irrefl: "\<And>M :: 'a multiset. \<not> ?less M M" |
|
43 |
proof |
|
44 |
fix M :: "'a multiset" |
|
45 |
have "trans {(x'::'a, x). x' < x}" |
|
46 |
by (rule transI) simp |
|
47 |
moreover |
|
48 |
assume "(M, M) \<in> mult {(x, y). x < y}" |
|
49 |
ultimately have "\<exists>I J K. M = I + J \<and> M = I + K |
|
60495 | 50 |
\<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset J. (k, j) \<in> {(x, y). x < y})" |
59813 | 51 |
by (rule mult_implies_one_step) |
52 |
then obtain I J K where "M = I + J" and "M = I + K" |
|
60495 | 53 |
and "J \<noteq> {#}" and "(\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset J. (k, j) \<in> {(x, y). x < y})" by blast |
54 |
then have aux1: "K \<noteq> {#}" and aux2: "\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset K. k < j" by auto |
|
55 |
have "finite (set_mset K)" by simp |
|
59813 | 56 |
moreover note aux2 |
60495 | 57 |
ultimately have "set_mset K = {}" |
59813 | 58 |
by (induct rule: finite_induct) |
59 |
(simp, metis (mono_tags) insert_absorb insert_iff insert_not_empty less_irrefl less_trans) |
|
60 |
with aux1 show False by simp |
|
61 |
qed |
|
62 |
have trans: "\<And>K M N :: 'a multiset. ?less K M \<Longrightarrow> ?less M N \<Longrightarrow> ?less K N" |
|
63 |
unfolding mult_def by (blast intro: trancl_trans) |
|
64 |
show "class.order ?le ?less" |
|
60679 | 65 |
by standard (auto simp add: le_multiset_def irrefl dest: trans) |
59813 | 66 |
qed |
67 |
||
60500 | 68 |
text \<open>The Dershowitz--Manna ordering:\<close> |
59813 | 69 |
|
70 |
definition less_multiset\<^sub>D\<^sub>M where |
|
71 |
"less_multiset\<^sub>D\<^sub>M M N \<longleftrightarrow> |
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
72 |
(\<exists>X Y. X \<noteq> {#} \<and> X \<le># N \<and> M = (N - X) + Y \<and> (\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)))" |
59813 | 73 |
|
74 |
||
60500 | 75 |
text \<open>The Huet--Oppen ordering:\<close> |
59813 | 76 |
|
77 |
definition less_multiset\<^sub>H\<^sub>O where |
|
78 |
"less_multiset\<^sub>H\<^sub>O M N \<longleftrightarrow> M \<noteq> N \<and> (\<forall>y. count N y < count M y \<longrightarrow> (\<exists>x. y < x \<and> count M x < count N x))" |
|
79 |
||
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
80 |
lemma mult_imp_less_multiset\<^sub>H\<^sub>O: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
81 |
"(M, N) \<in> mult {(x, y). x < y} \<Longrightarrow> less_multiset\<^sub>H\<^sub>O M N" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
82 |
proof (unfold mult_def, induct rule: trancl_induct) |
59813 | 83 |
case (base P) |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
84 |
then show ?case |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
85 |
by (auto elim!: mult1_lessE simp add: count_eq_zero_iff less_multiset\<^sub>H\<^sub>O_def split: if_splits dest!: Suc_lessD) |
59813 | 86 |
next |
87 |
case (step N P) |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
88 |
from step(3) have "M \<noteq> N" and |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
89 |
**: "\<And>y. count N y < count M y \<Longrightarrow> (\<exists>x>y. count M x < count N x)" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
90 |
by (simp_all add: less_multiset\<^sub>H\<^sub>O_def) |
59813 | 91 |
from step(2) obtain M0 a K where |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
92 |
*: "P = M0 + {#a#}" "N = M0 + K" "a \<notin># K" "\<And>b. b \<in># K \<Longrightarrow> b < a" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
93 |
by (blast elim: mult1_lessE) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
94 |
from \<open>M \<noteq> N\<close> ** *(1,2,3) have "M \<noteq> P" by (force dest: *(4) split: if_splits) |
59813 | 95 |
moreover |
96 |
{ assume "count P a \<le> count M a" |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
97 |
with \<open>a \<notin># K\<close> have "count N a < count M a" unfolding *(1,2) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
98 |
by (auto simp add: not_in_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
99 |
with ** obtain z where z: "z > a" "count M z < count N z" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
100 |
by blast |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
101 |
with * have "count N z \<le> count P z" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
102 |
by (force simp add: not_in_iff) |
59813 | 103 |
with z have "\<exists>z > a. count M z < count P z" by auto |
104 |
} note count_a = this |
|
105 |
{ fix y |
|
106 |
assume count_y: "count P y < count M y" |
|
107 |
have "\<exists>x>y. count M x < count P x" |
|
108 |
proof (cases "y = a") |
|
109 |
case True |
|
110 |
with count_y count_a show ?thesis by auto |
|
111 |
next |
|
112 |
case False |
|
113 |
show ?thesis |
|
114 |
proof (cases "y \<in># K") |
|
115 |
case True |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
116 |
with *(4) have "y < a" by simp |
59813 | 117 |
then show ?thesis by (cases "count P a \<le> count M a") (auto dest: count_a intro: less_trans) |
118 |
next |
|
119 |
case False |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
120 |
with \<open>y \<noteq> a\<close> have "count P y = count N y" unfolding *(1,2) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
121 |
by (simp add: not_in_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
122 |
with count_y ** obtain z where z: "z > y" "count M z < count N z" by auto |
59813 | 123 |
show ?thesis |
124 |
proof (cases "z \<in># K") |
|
125 |
case True |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
126 |
with *(4) have "z < a" by simp |
59813 | 127 |
with z(1) show ?thesis |
128 |
by (cases "count P a \<le> count M a") (auto dest!: count_a intro: less_trans) |
|
129 |
next |
|
130 |
case False |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
131 |
with \<open>a \<notin># K\<close> have "count N z \<le> count P z" unfolding * |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
132 |
by (auto simp add: not_in_iff) |
59813 | 133 |
with z show ?thesis by auto |
134 |
qed |
|
135 |
qed |
|
136 |
qed |
|
137 |
} |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
138 |
ultimately show ?case unfolding less_multiset\<^sub>H\<^sub>O_def by blast |
59813 | 139 |
qed |
140 |
||
141 |
lemma less_multiset\<^sub>D\<^sub>M_imp_mult: |
|
142 |
"less_multiset\<^sub>D\<^sub>M M N \<Longrightarrow> (M, N) \<in> mult {(x, y). x < y}" |
|
143 |
proof - |
|
144 |
assume "less_multiset\<^sub>D\<^sub>M M N" |
|
145 |
then obtain X Y where |
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
146 |
"X \<noteq> {#}" and "X \<le># N" and "M = N - X + Y" and "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)" |
59813 | 147 |
unfolding less_multiset\<^sub>D\<^sub>M_def by blast |
148 |
then have "(N - X + Y, N - X + X) \<in> mult {(x, y). x < y}" |
|
149 |
by (intro one_step_implies_mult) (auto simp: Bex_def trans_def) |
|
60500 | 150 |
with \<open>M = N - X + Y\<close> \<open>X \<le># N\<close> show "(M, N) \<in> mult {(x, y). x < y}" |
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
151 |
by (metis subset_mset.diff_add) |
59813 | 152 |
qed |
153 |
||
154 |
lemma less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M: "less_multiset\<^sub>H\<^sub>O M N \<Longrightarrow> less_multiset\<^sub>D\<^sub>M M N" |
|
155 |
unfolding less_multiset\<^sub>D\<^sub>M_def |
|
156 |
proof (intro iffI exI conjI) |
|
157 |
assume "less_multiset\<^sub>H\<^sub>O M N" |
|
158 |
then obtain z where z: "count M z < count N z" |
|
159 |
unfolding less_multiset\<^sub>H\<^sub>O_def by (auto simp: multiset_eq_iff nat_neq_iff) |
|
160 |
def X \<equiv> "N - M" |
|
161 |
def Y \<equiv> "M - N" |
|
162 |
from z show "X \<noteq> {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq) |
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
163 |
from z show "X \<le># N" unfolding X_def by auto |
59813 | 164 |
show "M = (N - X) + Y" unfolding X_def Y_def multiset_eq_iff count_union count_diff by force |
165 |
show "\<forall>k. k \<in># Y \<longrightarrow> (\<exists>a. a \<in># X \<and> k < a)" |
|
166 |
proof (intro allI impI) |
|
167 |
fix k |
|
168 |
assume "k \<in># Y" |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
169 |
then have "count N k < count M k" unfolding Y_def |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
170 |
by (auto simp add: in_diff_count) |
60500 | 171 |
with \<open>less_multiset\<^sub>H\<^sub>O M N\<close> obtain a where "k < a" and "count M a < count N a" |
59813 | 172 |
unfolding less_multiset\<^sub>H\<^sub>O_def by blast |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
173 |
then show "\<exists>a. a \<in># X \<and> k < a" unfolding X_def |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
174 |
by (auto simp add: in_diff_count) |
59813 | 175 |
qed |
176 |
qed |
|
177 |
||
178 |
lemma mult_less_multiset\<^sub>D\<^sub>M: "(M, N) \<in> mult {(x, y). x < y} \<longleftrightarrow> less_multiset\<^sub>D\<^sub>M M N" |
|
179 |
by (metis less_multiset\<^sub>D\<^sub>M_imp_mult less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M mult_imp_less_multiset\<^sub>H\<^sub>O) |
|
180 |
||
181 |
lemma mult_less_multiset\<^sub>H\<^sub>O: "(M, N) \<in> mult {(x, y). x < y} \<longleftrightarrow> less_multiset\<^sub>H\<^sub>O M N" |
|
182 |
by (metis less_multiset\<^sub>D\<^sub>M_imp_mult less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M mult_imp_less_multiset\<^sub>H\<^sub>O) |
|
183 |
||
184 |
lemmas mult\<^sub>D\<^sub>M = mult_less_multiset\<^sub>D\<^sub>M[unfolded less_multiset\<^sub>D\<^sub>M_def] |
|
185 |
lemmas mult\<^sub>H\<^sub>O = mult_less_multiset\<^sub>H\<^sub>O[unfolded less_multiset\<^sub>H\<^sub>O_def] |
|
186 |
||
187 |
end |
|
188 |
||
189 |
context linorder |
|
190 |
begin |
|
191 |
||
61076 | 192 |
lemma total_le: "total {(a :: 'a, b). a \<le> b}" |
59813 | 193 |
unfolding total_on_def by auto |
194 |
||
61076 | 195 |
lemma total_less: "total {(a :: 'a, b). a < b}" |
59813 | 196 |
unfolding total_on_def by auto |
197 |
||
198 |
lemma linorder_mult: "class.linorder |
|
199 |
(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y} \<or> M = N) |
|
200 |
(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y})" |
|
201 |
proof - |
|
202 |
interpret o: order |
|
203 |
"(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y} \<or> M = N)" |
|
204 |
"(\<lambda>M N. (M, N) \<in> mult {(x, y). x < y})" |
|
205 |
by (rule order_mult) |
|
206 |
show ?thesis by unfold_locales (auto 0 3 simp: mult\<^sub>H\<^sub>O not_less_iff_gr_or_eq) |
|
207 |
qed |
|
208 |
||
209 |
end |
|
210 |
||
211 |
lemma less_multiset_less_multiset\<^sub>H\<^sub>O: |
|
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
212 |
"M #\<subset># N \<longleftrightarrow> less_multiset\<^sub>H\<^sub>O M N" |
59813 | 213 |
unfolding less_multiset_def mult\<^sub>H\<^sub>O less_multiset\<^sub>H\<^sub>O_def .. |
214 |
||
215 |
lemmas less_multiset\<^sub>D\<^sub>M = mult\<^sub>D\<^sub>M[folded less_multiset_def] |
|
216 |
lemmas less_multiset\<^sub>H\<^sub>O = mult\<^sub>H\<^sub>O[folded less_multiset_def] |
|
217 |
||
218 |
lemma le_multiset\<^sub>H\<^sub>O: |
|
61076 | 219 |
fixes M N :: "('a :: linorder) multiset" |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
220 |
shows "M #\<subseteq># N \<longleftrightarrow> (\<forall>y. count N y < count M y \<longrightarrow> (\<exists>x. y < x \<and> count M x < count N x))" |
59813 | 221 |
by (auto simp: le_multiset_def less_multiset\<^sub>H\<^sub>O) |
222 |
||
61076 | 223 |
lemma wf_less_multiset: "wf {(M :: ('a :: wellorder) multiset, N). M #\<subset># N}" |
59813 | 224 |
unfolding less_multiset_def by (auto intro: wf_mult wf) |
225 |
||
226 |
lemma order_multiset: "class.order |
|
61076 | 227 |
(le_multiset :: ('a :: order) multiset \<Rightarrow> ('a :: order) multiset \<Rightarrow> bool) |
228 |
(less_multiset :: ('a :: order) multiset \<Rightarrow> ('a :: order) multiset \<Rightarrow> bool)" |
|
59813 | 229 |
by unfold_locales |
230 |
||
231 |
lemma linorder_multiset: "class.linorder |
|
61076 | 232 |
(le_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool) |
233 |
(less_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool)" |
|
59813 | 234 |
by unfold_locales (fastforce simp add: less_multiset\<^sub>H\<^sub>O le_multiset_def not_less_iff_gr_or_eq) |
235 |
||
236 |
interpretation multiset_linorder: linorder |
|
61076 | 237 |
"le_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool" |
238 |
"less_multiset :: ('a :: linorder) multiset \<Rightarrow> ('a :: linorder) multiset \<Rightarrow> bool" |
|
59813 | 239 |
by (rule linorder_multiset) |
240 |
||
241 |
interpretation multiset_wellorder: wellorder |
|
61076 | 242 |
"le_multiset :: ('a :: wellorder) multiset \<Rightarrow> ('a :: wellorder) multiset \<Rightarrow> bool" |
243 |
"less_multiset :: ('a :: wellorder) multiset \<Rightarrow> ('a :: wellorder) multiset \<Rightarrow> bool" |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61076
diff
changeset
|
244 |
by unfold_locales (blast intro: wf_less_multiset [unfolded wf_def, simplified, rule_format]) |
59813 | 245 |
|
246 |
lemma le_multiset_total: |
|
61076 | 247 |
fixes M N :: "('a :: linorder) multiset" |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
248 |
shows "\<not> M #\<subseteq># N \<Longrightarrow> N #\<subseteq># M" |
59813 | 249 |
by (metis multiset_linorder.le_cases) |
250 |
||
251 |
lemma less_eq_imp_le_multiset: |
|
61076 | 252 |
fixes M N :: "('a :: linorder) multiset" |
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
253 |
shows "M \<le># N \<Longrightarrow> M #\<subseteq># N" |
59813 | 254 |
unfolding le_multiset_def less_multiset\<^sub>H\<^sub>O |
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
255 |
by (simp add: less_le_not_le subseteq_mset_def) |
59813 | 256 |
|
257 |
lemma less_multiset_right_total: |
|
61076 | 258 |
fixes M :: "('a :: linorder) multiset" |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
259 |
shows "M #\<subset># M + {#undefined#}" |
59813 | 260 |
unfolding le_multiset_def less_multiset\<^sub>H\<^sub>O by simp |
261 |
||
262 |
lemma le_multiset_empty_left[simp]: |
|
61076 | 263 |
fixes M :: "('a :: linorder) multiset" |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
264 |
shows "{#} #\<subseteq># M" |
59813 | 265 |
by (simp add: less_eq_imp_le_multiset) |
266 |
||
267 |
lemma le_multiset_empty_right[simp]: |
|
61076 | 268 |
fixes M :: "('a :: linorder) multiset" |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
269 |
shows "M \<noteq> {#} \<Longrightarrow> \<not> M #\<subseteq># {#}" |
59813 | 270 |
by (metis le_multiset_empty_left multiset_order.antisym) |
271 |
||
272 |
lemma less_multiset_empty_left[simp]: |
|
61076 | 273 |
fixes M :: "('a :: linorder) multiset" |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
274 |
shows "M \<noteq> {#} \<Longrightarrow> {#} #\<subset># M" |
59813 | 275 |
by (simp add: less_multiset\<^sub>H\<^sub>O) |
276 |
||
277 |
lemma less_multiset_empty_right[simp]: |
|
61076 | 278 |
fixes M :: "('a :: linorder) multiset" |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
279 |
shows "\<not> M #\<subset># {#}" |
59813 | 280 |
using le_empty less_multiset\<^sub>D\<^sub>M by blast |
281 |
||
282 |
lemma |
|
61076 | 283 |
fixes M N :: "('a :: linorder) multiset" |
59813 | 284 |
shows |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
285 |
le_multiset_plus_left[simp]: "N #\<subseteq># (M + N)" and |
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
286 |
le_multiset_plus_right[simp]: "M #\<subseteq># (M + N)" |
59813 | 287 |
using [[metis_verbose = false]] by (metis less_eq_imp_le_multiset mset_le_add_left add.commute)+ |
288 |
||
289 |
lemma |
|
61076 | 290 |
fixes M N :: "('a :: linorder) multiset" |
59813 | 291 |
shows |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
292 |
less_multiset_plus_plus_left_iff[simp]: "M + N #\<subset># M' + N \<longleftrightarrow> M #\<subset># M'" and |
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
293 |
less_multiset_plus_plus_right_iff[simp]: "M + N #\<subset># M + N' \<longleftrightarrow> N #\<subset># N'" |
59813 | 294 |
unfolding less_multiset\<^sub>H\<^sub>O by auto |
295 |
||
296 |
lemma add_eq_self_empty_iff: "M + N = M \<longleftrightarrow> N = {#}" |
|
297 |
by (metis add.commute add_diff_cancel_right' monoid_add_class.add.left_neutral) |
|
298 |
||
299 |
lemma |
|
61076 | 300 |
fixes M N :: "('a :: linorder) multiset" |
59813 | 301 |
shows |
59958
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
302 |
less_multiset_plus_left_nonempty[simp]: "M \<noteq> {#} \<Longrightarrow> N #\<subset># M + N" and |
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
blanchet
parents:
59817
diff
changeset
|
303 |
less_multiset_plus_right_nonempty[simp]: "N \<noteq> {#} \<Longrightarrow> M #\<subset># M + N" |
59813 | 304 |
using [[metis_verbose = false]] |
305 |
by (metis add.right_neutral less_multiset_empty_left less_multiset_plus_plus_right_iff |
|
306 |
add.commute)+ |
|
307 |
||
61076 | 308 |
lemma ex_gt_imp_less_multiset: "(\<exists>y :: 'a :: linorder. y \<in># N \<and> (\<forall>x. x \<in># M \<longrightarrow> x < y)) \<Longrightarrow> M #\<subset># N" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
309 |
unfolding less_multiset\<^sub>H\<^sub>O |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
310 |
by (metis count_eq_zero_iff count_greater_zero_iff less_le_not_le) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
311 |
|
59813 | 312 |
lemma ex_gt_count_imp_less_multiset: |
61076 | 313 |
"(\<forall>y :: 'a :: linorder. y \<in># M + N \<longrightarrow> y \<le> x) \<Longrightarrow> count M x < count N x \<Longrightarrow> M #\<subset># N" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
314 |
unfolding less_multiset\<^sub>H\<^sub>O |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
61424
diff
changeset
|
315 |
by (metis add_gr_0 count_union mem_Collect_eq not_gr0 not_le not_less_iff_gr_or_eq set_mset_def) |
59813 | 316 |
|
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
317 |
lemma union_less_diff_plus: "P \<le># M \<Longrightarrow> N #\<subset># P \<Longrightarrow> M - P + N #\<subset># M" |
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59958
diff
changeset
|
318 |
by (drule subset_mset.diff_add[symmetric]) (metis union_less_mono2) |
59813 | 319 |
|
320 |
end |