author | blanchet |
Wed, 06 Mar 2024 10:39:45 +0100 | |
changeset 79800 | abb5e57c92a7 |
parent 79575 | b21d8401f0ca |
child 79971 | 033f90dc441d |
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 |
||
79800 | 246 |
lemma count_gt_imp_in_mset: "count M x > n \<Longrightarrow> x \<in># M" |
247 |
using count_greater_zero_iff by fastforce |
|
248 |
||
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
249 |
|
60500 | 250 |
subsubsection \<open>Union\<close> |
10249 | 251 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
252 |
lemma count_union [simp]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
253 |
"count (M + N) a = count M a + count N a" |
47429
ec64d94cbf9c
multiset operations are defined with lift_definitions;
bulwahn
parents:
47308
diff
changeset
|
254 |
by (simp add: plus_multiset.rep_eq) |
10249 | 255 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
256 |
lemma set_mset_union [simp]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
257 |
"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
|
258 |
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
|
259 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
260 |
lemma union_mset_add_mset_left [simp]: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
261 |
"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
|
262 |
by (auto simp: multiset_eq_iff) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
263 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
264 |
lemma union_mset_add_mset_right [simp]: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
265 |
"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
|
266 |
by (auto simp: multiset_eq_iff) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
267 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
268 |
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
|
269 |
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
|
270 |
|
10249 | 271 |
|
60500 | 272 |
subsubsection \<open>Difference\<close> |
10249 | 273 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
274 |
instance multiset :: (type) comm_monoid_diff |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
275 |
by standard (transfer; simp add: fun_eq_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
276 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
277 |
lemma count_diff [simp]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
278 |
"count (M - N) a = count M a - count N a" |
47429
ec64d94cbf9c
multiset operations are defined with lift_definitions;
bulwahn
parents:
47308
diff
changeset
|
279 |
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
|
280 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
281 |
lemma add_mset_diff_bothsides: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
282 |
\<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
|
283 |
by (auto simp: multiset_eq_iff) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
284 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
285 |
lemma in_diff_count: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
286 |
"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
|
287 |
by (simp add: set_mset_def) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
288 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
289 |
lemma count_in_diffI: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
290 |
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
|
291 |
shows "x \<in># M - N" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
292 |
proof (rule ccontr) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
293 |
assume "x \<notin># M - N" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
294 |
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
|
295 |
by (simp add: in_diff_count not_less) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
296 |
with assms show False by auto |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
297 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
298 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
299 |
lemma in_diff_countE: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
300 |
assumes "x \<in># M - N" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
301 |
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
|
302 |
proof - |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
303 |
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
|
304 |
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
|
305 |
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
|
306 |
using less_iff_Suc_add by auto |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
307 |
with that show thesis . |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
308 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
309 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
310 |
lemma in_diffD: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
311 |
assumes "a \<in># M - N" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
312 |
shows "a \<in># M" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
313 |
proof - |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
314 |
have "0 \<le> count N a" by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
315 |
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
|
316 |
by (simp add: in_diff_count) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
317 |
finally show ?thesis by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
318 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
319 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
320 |
lemma set_mset_diff: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
321 |
"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
|
322 |
by (simp add: set_mset_def) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
323 |
|
17161 | 324 |
lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}" |
52289 | 325 |
by rule (fact Groups.diff_zero, fact Groups.zero_diff) |
36903 | 326 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
327 |
lemma diff_cancel: "A - A = {#}" |
52289 | 328 |
by (fact Groups.diff_cancel) |
10249 | 329 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
330 |
lemma diff_union_cancelR: "M + N - N = (M::'a multiset)" |
52289 | 331 |
by (fact add_diff_cancel_right') |
10249 | 332 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
333 |
lemma diff_union_cancelL: "N + M - N = (M::'a multiset)" |
52289 | 334 |
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
|
335 |
|
52289 | 336 |
lemma diff_right_commute: |
60606 | 337 |
fixes M N Q :: "'a multiset" |
338 |
shows "M - N - Q = M - Q - N" |
|
52289 | 339 |
by (fact diff_right_commute) |
340 |
||
341 |
lemma diff_add: |
|
60606 | 342 |
fixes M N Q :: "'a multiset" |
343 |
shows "M - (N + Q) = M - N - Q" |
|
52289 | 344 |
by (rule sym) (fact diff_diff_add) |
58425 | 345 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
346 |
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
|
347 |
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
|
348 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
349 |
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
|
350 |
by simp |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
351 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
352 |
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
|
353 |
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
|
354 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
355 |
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
|
356 |
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
|
357 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
358 |
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
|
359 |
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
|
360 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
361 |
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
|
362 |
by (rule diff_diff_add) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
363 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
364 |
lemma diff_union_single_conv: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
365 |
"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
|
366 |
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
|
367 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
368 |
lemma mset_add [elim?]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
369 |
assumes "a \<in># A" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
370 |
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
|
371 |
proof - |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
372 |
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
|
373 |
by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
374 |
with that show thesis . |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
375 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
376 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
377 |
lemma union_iff: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
378 |
"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
|
379 |
by auto |
26143
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
bulwahn
parents:
26033
diff
changeset
|
380 |
|
77987
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
desharna
parents:
77832
diff
changeset
|
381 |
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
|
382 |
"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
|
383 |
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
|
384 |
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
|
385 |
|
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
desharna
parents:
77832
diff
changeset
|
386 |
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
|
387 |
"(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
|
388 |
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
|
389 |
|
10249 | 390 |
|
66425 | 391 |
subsubsection \<open>Min and Max\<close> |
392 |
||
393 |
abbreviation Min_mset :: "'a::linorder multiset \<Rightarrow> 'a" where |
|
394 |
"Min_mset m \<equiv> Min (set_mset m)" |
|
395 |
||
396 |
abbreviation Max_mset :: "'a::linorder multiset \<Rightarrow> 'a" where |
|
397 |
"Max_mset m \<equiv> Max (set_mset m)" |
|
398 |
||
79800 | 399 |
lemma |
400 |
Min_in_mset: "M \<noteq> {#} \<Longrightarrow> Min_mset M \<in># M" and |
|
401 |
Max_in_mset: "M \<noteq> {#} \<Longrightarrow> Max_mset M \<in># M" |
|
402 |
by simp+ |
|
403 |
||
66425 | 404 |
|
60500 | 405 |
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
|
406 |
|
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 |
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
|
408 |
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
|
409 |
|
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 |
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
|
411 |
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
|
412 |
|
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
|
413 |
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
|
414 |
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
|
415 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
416 |
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
|
417 |
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
|
418 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
419 |
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
|
420 |
by (auto simp: multiset_eq_iff) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
421 |
|
60606 | 422 |
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
|
423 |
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
|
424 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
425 |
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
|
426 |
by auto |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
427 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
428 |
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
|
429 |
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
|
430 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
431 |
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
|
432 |
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
|
433 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
434 |
lemma add_mset_remove_trivial_If: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
435 |
"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
|
436 |
by (simp add: diff_single_trivial) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
437 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
438 |
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
|
439 |
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
|
440 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
441 |
lemma union_is_single: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
442 |
"M + N = {#a#} \<longleftrightarrow> M = {#a#} \<and> N = {#} \<or> M = {#} \<and> N = {#a#}" |
60606 | 443 |
(is "?lhs = ?rhs") |
46730 | 444 |
proof |
60606 | 445 |
show ?lhs if ?rhs using that by auto |
446 |
show ?rhs if ?lhs |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
447 |
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
|
448 |
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
|
449 |
|
60606 | 450 |
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
|
451 |
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
|
452 |
|
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
|
453 |
lemma add_eq_conv_diff: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
454 |
"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 | 455 |
(is "?lhs \<longleftrightarrow> ?rhs") |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
44339
diff
changeset
|
456 |
(* 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
|
457 |
proof |
60606 | 458 |
show ?lhs if ?rhs |
459 |
using that |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
460 |
by (auto simp add: add_mset_commute[of a b]) |
60606 | 461 |
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
|
462 |
proof (cases "a = b") |
60500 | 463 |
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
|
464 |
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
|
465 |
case False |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
466 |
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
|
467 |
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
|
468 |
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
|
469 |
moreover note False |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
470 |
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
|
471 |
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
|
472 |
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
|
473 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
474 |
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
|
475 |
by (auto simp: add_eq_conv_diff) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
476 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
477 |
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
|
478 |
by (auto simp: add_eq_conv_diff) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
479 |
|
58425 | 480 |
lemma insert_noteq_member: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
481 |
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
|
482 |
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
|
483 |
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
|
484 |
proof - |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
485 |
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
|
486 |
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
|
487 |
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
|
488 |
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
|
489 |
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
|
490 |
|
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
|
491 |
lemma add_eq_conv_ex: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
492 |
"(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
|
493 |
(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
|
494 |
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
|
495 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
496 |
lemma multi_member_split: "x \<in># M \<Longrightarrow> \<exists>A. M = add_mset x A" |
60678 | 497 |
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
|
498 |
|
58425 | 499 |
lemma multiset_add_sub_el_shuffle: |
60606 | 500 |
assumes "c \<in># B" |
501 |
and "b \<noteq> c" |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
502 |
shows "add_mset b (B - {#c#}) = add_mset b B - {#c#}" |
58098 | 503 |
proof - |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
504 |
from \<open>c \<in># B\<close> obtain A where B: "B = add_mset c A" |
58098 | 505 |
by (blast dest: multi_member_split) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
506 |
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
|
507 |
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
|
508 |
by (simp add: \<open>b \<noteq> c\<close>) |
58098 | 509 |
then show ?thesis using B by simp |
510 |
qed |
|
511 |
||
64418 | 512 |
lemma add_mset_eq_singleton_iff[iff]: |
513 |
"add_mset x M = {#y#} \<longleftrightarrow> M = {#} \<and> x = y" |
|
514 |
by auto |
|
515 |
||
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
|
516 |
|
60500 | 517 |
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
|
518 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
519 |
definition subseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subseteq>#" 50) |
65466 | 520 |
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
|
521 |
|
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
522 |
definition subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subset>#" 50) |
65466 | 523 |
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
|
524 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
525 |
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
|
526 |
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
|
527 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
528 |
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
|
529 |
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
|
530 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
531 |
notation (input) |
62208
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
blanchet
parents:
62082
diff
changeset
|
532 |
subseteq_mset (infix "\<le>#" 50) and |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
533 |
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
|
534 |
|
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
535 |
notation (ASCII) |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
536 |
subseteq_mset (infix "<=#" 50) and |
62208
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
blanchet
parents:
62082
diff
changeset
|
537 |
subset_mset (infix "<#" 50) and |
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
blanchet
parents:
62082
diff
changeset
|
538 |
supseteq_mset (infix ">=#" 50) and |
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
blanchet
parents:
62082
diff
changeset
|
539 |
supset_mset (infix ">#" 50) |
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59999
diff
changeset
|
540 |
|
73411 | 541 |
global_interpretation subset_mset: ordering \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> |
542 |
by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order.trans order.antisym) |
|
543 |
||
73451 | 544 |
interpretation subset_mset: ordered_ab_semigroup_add_imp_le \<open>(+)\<close> \<open>(-)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> |
60678 | 545 |
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
|
546 |
\<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
|
547 |
|
67398 | 548 |
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
|
549 |
by standard |
64585
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
haftmann
parents:
64531
diff
changeset
|
550 |
\<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
|
551 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
552 |
lemma mset_subset_eqI: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
553 |
"(\<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
|
554 |
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
|
555 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
556 |
lemma mset_subset_eq_count: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
557 |
"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
|
558 |
by (simp add: subseteq_mset_def) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
559 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
560 |
lemma mset_subset_eq_exists_conv: "(A::'a multiset) \<subseteq># B \<longleftrightarrow> (\<exists>C. B = A + C)" |
60678 | 561 |
unfolding subseteq_mset_def |
562 |
apply (rule iffI) |
|
563 |
apply (rule exI [where x = "B - A"]) |
|
564 |
apply (auto intro: multiset_eq_iff [THEN iffD2]) |
|
565 |
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
|
566 |
|
67398 | 567 |
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
|
568 |
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
|
569 |
\<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
|
570 |
|
64017
6e7bf7678518
more multiset simp rules
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63924
diff
changeset
|
571 |
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
|
572 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
573 |
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
|
574 |
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
|
575 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
576 |
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
|
577 |
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
|
578 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
579 |
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
|
580 |
by (fact subset_mset.add_mono) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
581 |
|
63560
3e3097ac37d1
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63547
diff
changeset
|
582 |
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
|
583 |
by simp |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
584 |
|
63560
3e3097ac37d1
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63547
diff
changeset
|
585 |
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
|
586 |
by simp |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
587 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
588 |
lemma single_subset_iff [simp]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
589 |
"{#a#} \<subseteq># M \<longleftrightarrow> a \<in># M" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
590 |
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
|
591 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
592 |
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
|
593 |
by simp |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
594 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
595 |
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
|
596 |
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
|
597 |
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
|
598 |
|
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
599 |
lemma multiset_diff_union_assoc: |
60606 | 600 |
fixes A B C D :: "'a multiset" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
601 |
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
|
602 |
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
|
603 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
604 |
lemma mset_subset_eq_multiset_union_diff_commute: |
60606 | 605 |
fixes A B C D :: "'a multiset" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
606 |
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
|
607 |
by (fact subset_mset.add_diff_assoc2) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
608 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
609 |
lemma diff_subset_eq_self[simp]: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
610 |
"(M::'a multiset) - N \<subseteq># M" |
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 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
613 |
lemma mset_subset_eqD: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
614 |
assumes "A \<subseteq># B" and "x \<in># A" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
615 |
shows "x \<in># B" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
616 |
proof - |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
617 |
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
|
618 |
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
|
619 |
by (simp add: subseteq_mset_def) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
620 |
finally show ?thesis by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
621 |
qed |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
622 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
623 |
lemma mset_subsetD: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
624 |
"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
|
625 |
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
|
626 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
627 |
lemma set_mset_mono: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
628 |
"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
|
629 |
by (metis mset_subset_eqD subsetI) |
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
630 |
|
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
631 |
lemma mset_subset_eq_insertD: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
632 |
"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
|
633 |
apply (rule conjI) |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
634 |
apply (simp add: mset_subset_eqD) |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
635 |
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
|
636 |
apply safe |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
637 |
apply (erule_tac x = a in allE) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
638 |
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
|
639 |
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
|
640 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
641 |
lemma mset_subset_insertD: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
642 |
"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
|
643 |
by (rule mset_subset_eq_insertD) simp |
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
644 |
|
63831 | 645 |
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
|
646 |
by (simp only: subset_mset.not_less_zero) |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
647 |
|
64587 | 648 |
lemma empty_subset_add_mset[simp]: "{#} \<subset># add_mset x M" |
649 |
by (auto intro: subset_mset.gr_zeroI) |
|
63831 | 650 |
|
63795
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
651 |
lemma empty_le: "{#} \<subseteq># A" |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
652 |
by (fact subset_mset.zero_le) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
653 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
654 |
lemma insert_subset_eq_iff: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
655 |
"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
|
656 |
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
|
657 |
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
|
658 |
apply (rule ccontr) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
659 |
apply (auto simp add: not_in_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
660 |
done |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
661 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
662 |
lemma insert_union_subset_iff: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
663 |
"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
|
664 |
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
|
665 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
666 |
lemma subset_eq_diff_conv: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
667 |
"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
|
668 |
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
|
669 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
670 |
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
|
671 |
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
|
672 |
|
64076 | 673 |
lemma multi_psub_self: "A \<subset># A = False" |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
674 |
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
|
675 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
676 |
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
|
677 |
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
|
678 |
by (fact subset_mset.add_less_cancel_right) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
679 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
680 |
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
|
681 |
by (auto simp: subset_mset_def elim: mset_add) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
682 |
|
64077 | 683 |
lemma Diff_eq_empty_iff_mset: "A - B = {#} \<longleftrightarrow> A \<subseteq># B" |
684 |
by (auto simp: multiset_eq_iff subseteq_mset_def) |
|
685 |
||
64418 | 686 |
lemma add_mset_subseteq_single_iff[iff]: "add_mset a M \<subseteq># {#b#} \<longleftrightarrow> M = {#} \<and> a = b" |
687 |
proof |
|
688 |
assume A: "add_mset a M \<subseteq># {#b#}" |
|
689 |
then have \<open>a = b\<close> |
|
690 |
by (auto dest: mset_subset_eq_insertD) |
|
691 |
then show "M={#} \<and> a=b" |
|
692 |
using A by (simp add: mset_subset_eq_add_mset_cancel) |
|
693 |
qed simp |
|
694 |
||
79800 | 695 |
lemma nonempty_subseteq_mset_eq_single: "M \<noteq> {#} \<Longrightarrow> M \<subseteq># {#x#} \<Longrightarrow> M = {#x#}" |
696 |
by (cases M) (metis single_is_union subset_mset.less_eqE) |
|
697 |
||
698 |
lemma nonempty_subseteq_mset_iff_single: "(M \<noteq> {#} \<and> M \<subseteq># {#x#} \<and> P) \<longleftrightarrow> M = {#x#} \<and> P" |
|
699 |
by (cases M) (metis empty_not_add_mset nonempty_subseteq_mset_eq_single subset_mset.order_refl) |
|
700 |
||
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
701 |
|
64076 | 702 |
subsubsection \<open>Intersection and bounded union\<close> |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
703 |
|
73393 | 704 |
definition inter_mset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> (infixl \<open>\<inter>#\<close> 70) |
705 |
where \<open>A \<inter># B = A - (A - B)\<close> |
|
706 |
||
73451 | 707 |
lemma count_inter_mset [simp]: |
708 |
\<open>count (A \<inter># B) x = min (count A x) (count B x)\<close> |
|
709 |
by (simp add: inter_mset_def) |
|
710 |
||
711 |
(*global_interpretation subset_mset: semilattice_order \<open>(\<inter>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> |
|
712 |
by standard (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def min_def)*) |
|
713 |
||
73393 | 714 |
interpretation subset_mset: semilattice_inf \<open>(\<inter>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> |
73451 | 715 |
by standard (simp_all add: multiset_eq_iff subseteq_mset_def) |
716 |
\<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
|
717 |
|
73393 | 718 |
definition union_mset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> (infixl \<open>\<union>#\<close> 70) |
719 |
where \<open>A \<union># B = A + (B - A)\<close> |
|
720 |
||
73451 | 721 |
lemma count_union_mset [simp]: |
722 |
\<open>count (A \<union># B) x = max (count A x) (count B x)\<close> |
|
723 |
by (simp add: union_mset_def) |
|
724 |
||
725 |
global_interpretation subset_mset: semilattice_neutr_order \<open>(\<union>#)\<close> \<open>{#}\<close> \<open>(\<supseteq>#)\<close> \<open>(\<supset>#)\<close> |
|
726 |
apply standard |
|
727 |
apply (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def max_def) |
|
728 |
apply (auto simp add: le_antisym dest: sym) |
|
729 |
apply (metis nat_le_linear)+ |
|
730 |
done |
|
731 |
||
73393 | 732 |
interpretation subset_mset: semilattice_sup \<open>(\<union>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> |
64076 | 733 |
proof - |
734 |
have [simp]: "m \<le> n \<Longrightarrow> q \<le> n \<Longrightarrow> m + (q - m) \<le> n" for m n q :: nat |
|
735 |
by arith |
|
67398 | 736 |
show "class.semilattice_sup (\<union>#) (\<subseteq>#) (\<subset>#)" |
73393 | 737 |
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
|
738 |
qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> |
64076 | 739 |
|
67398 | 740 |
interpretation subset_mset: bounded_lattice_bot "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)" |
741 |
"(\<union>#)" "{#}" |
|
64076 | 742 |
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
|
743 |
\<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> |
64076 | 744 |
|
745 |
||
746 |
subsubsection \<open>Additional intersection facts\<close> |
|
747 |
||
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
748 |
lemma set_mset_inter [simp]: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
749 |
"set_mset (A \<inter># B) = set_mset A \<inter> set_mset B" |
73393 | 750 |
by (simp only: set_mset_def) auto |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
751 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
752 |
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
|
753 |
"M - M \<inter># N = M - N" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
754 |
by (simp add: multiset_eq_iff min_def) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
755 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
756 |
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
|
757 |
"M - N \<inter># M = M - N" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
758 |
by (simp add: multiset_eq_iff min_def) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
759 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
760 |
lemma multiset_inter_single[simp]: "a \<noteq> b \<Longrightarrow> {#a#} \<inter># {#b#} = {#}" |
46730 | 761 |
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
|
762 |
|
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
763 |
lemma multiset_union_diff_commute: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
764 |
assumes "B \<inter># C = {#}" |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
765 |
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
|
766 |
proof (rule multiset_eqI) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
767 |
fix x |
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
768 |
from assms have "min (count B x) (count C x) = 0" |
46730 | 769 |
by (auto simp add: multiset_eq_iff) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
770 |
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
|
771 |
unfolding min_def by (auto split: if_splits) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
772 |
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
|
773 |
by auto |
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
774 |
qed |
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
775 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
776 |
lemma disjunct_not_in: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
777 |
"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
|
778 |
proof |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
779 |
assume ?P |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
780 |
show ?Q |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
781 |
proof |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
782 |
fix a |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
783 |
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
|
784 |
by (simp add: multiset_eq_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
785 |
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
|
786 |
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
|
787 |
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
|
788 |
by (simp add: not_in_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
789 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
790 |
next |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
791 |
assume ?Q |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
792 |
show ?P |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
793 |
proof (rule multiset_eqI) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
794 |
fix a |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
795 |
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
|
796 |
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
|
797 |
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
|
798 |
by auto |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
799 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
800 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
801 |
|
64077 | 802 |
lemma inter_mset_empty_distrib_right: "A \<inter># (B + C) = {#} \<longleftrightarrow> A \<inter># B = {#} \<and> A \<inter># C = {#}" |
803 |
by (meson disjunct_not_in union_iff) |
|
804 |
||
805 |
lemma inter_mset_empty_distrib_left: "(A + B) \<inter># C = {#} \<longleftrightarrow> A \<inter># C = {#} \<and> B \<inter># C = {#}" |
|
806 |
by (meson disjunct_not_in union_iff) |
|
807 |
||
73393 | 808 |
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
|
809 |
"add_mset a A \<inter># add_mset a B = add_mset a (A \<inter># B)" |
73393 | 810 |
by (rule multiset_eqI) simp |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
811 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
812 |
lemma add_mset_disjoint [simp]: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
813 |
"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
|
814 |
"{#} = 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
|
815 |
by (auto simp: disjunct_not_in) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
816 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
817 |
lemma disjoint_add_mset [simp]: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
818 |
"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
|
819 |
"{#} = 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
|
820 |
by (auto simp: disjunct_not_in) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
821 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
822 |
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
|
823 |
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
|
824 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
825 |
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
|
826 |
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
|
827 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
828 |
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
|
829 |
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
|
830 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
831 |
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
|
832 |
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
|
833 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
834 |
lemma disjunct_set_mset_diff: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
835 |
assumes "M \<inter># N = {#}" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
836 |
shows "set_mset (M - N) = set_mset M" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
837 |
proof (rule set_eqI) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
838 |
fix a |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
839 |
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
|
840 |
by (simp add: disjunct_not_in) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
841 |
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
|
842 |
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
|
843 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
844 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
845 |
lemma at_most_one_mset_mset_diff: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
846 |
assumes "a \<notin># M - {#a#}" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
847 |
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
|
848 |
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
|
849 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
850 |
lemma more_than_one_mset_mset_diff: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
851 |
assumes "a \<in># M - {#a#}" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
852 |
shows "set_mset (M - {#a#}) = set_mset M" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
853 |
proof (rule set_eqI) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
854 |
fix b |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
855 |
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
|
856 |
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
|
857 |
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
|
858 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
859 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
860 |
lemma inter_iff: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
861 |
"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
|
862 |
by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
863 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
864 |
lemma inter_union_distrib_left: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
865 |
"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
|
866 |
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
|
867 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
868 |
lemma inter_union_distrib_right: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
869 |
"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
|
870 |
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
|
871 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
872 |
lemma inter_subset_eq_union: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
873 |
"A \<inter># B \<subseteq># A + B" |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
874 |
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
|
875 |
|
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
876 |
|
64076 | 877 |
subsubsection \<open>Additional bounded union facts\<close> |
63795
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
878 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
879 |
lemma set_mset_sup [simp]: |
73393 | 880 |
\<open>set_mset (A \<union># B) = set_mset A \<union> set_mset B\<close> |
881 |
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
|
882 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
883 |
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
|
884 |
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
|
885 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
886 |
lemma sup_union_left2: "x \<in># N \<Longrightarrow> (add_mset x M) \<union># N = add_mset x (M \<union># (N - {#x#}))" |
51623 | 887 |
by (simp add: multiset_eq_iff) |
888 |
||
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
889 |
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
|
890 |
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
|
891 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
892 |
lemma sup_union_right2: "x \<in># N \<Longrightarrow> N \<union># (add_mset x M) = add_mset x ((N - {#x#}) \<union># M)" |
51623 | 893 |
by (simp add: multiset_eq_iff) |
894 |
||
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
895 |
lemma sup_union_distrib_left: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
896 |
"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
|
897 |
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
|
898 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
899 |
lemma union_sup_distrib_right: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
900 |
"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
|
901 |
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
|
902 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
903 |
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
|
904 |
"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
|
905 |
by (auto simp add: multiset_eq_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
906 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
907 |
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
|
908 |
"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
|
909 |
by (auto simp add: multiset_eq_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
910 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
911 |
lemma add_mset_union: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
912 |
\<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
|
913 |
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
|
914 |
|
51623 | 915 |
|
63908
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
916 |
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
|
917 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
918 |
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
|
919 |
"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
|
920 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
921 |
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
|
922 |
unfolding replicate_mset_def by simp |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
923 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
924 |
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
|
925 |
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
|
926 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
927 |
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
|
928 |
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
|
929 |
|
73393 | 930 |
lift_definition repeat_mset :: \<open>nat \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> |
931 |
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
|
932 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
933 |
lemma count_repeat_mset [simp]: "count (repeat_mset i A) a = i * count A a" |
73393 | 934 |
by transfer rule |
935 |
||
936 |
lemma repeat_mset_0 [simp]: |
|
937 |
\<open>repeat_mset 0 M = {#}\<close> |
|
938 |
by transfer simp |
|
939 |
||
940 |
lemma repeat_mset_Suc [simp]: |
|
941 |
\<open>repeat_mset (Suc n) M = M + repeat_mset n M\<close> |
|
942 |
by transfer simp |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
943 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
944 |
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
|
945 |
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
|
946 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
947 |
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
|
948 |
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
|
949 |
|
63908
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
950 |
lemma left_add_mult_distrib_mset: |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
951 |
"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
|
952 |
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
|
953 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
954 |
lemma repeat_mset_distrib: |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
955 |
"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
|
956 |
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
|
957 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
958 |
lemma repeat_mset_distrib2[simp]: |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
959 |
"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
|
960 |
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
|
961 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
962 |
lemma repeat_mset_replicate_mset[simp]: |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
963 |
"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
|
964 |
by (auto simp: multiset_eq_iff) |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
965 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
966 |
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
|
967 |
"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
|
968 |
by (auto simp: multiset_eq_iff) |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
969 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
970 |
lemma repeat_mset_empty[simp]: "repeat_mset n {#} = {#}" |
73393 | 971 |
by transfer simp |
63908
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
972 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
973 |
|
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
974 |
subsubsection \<open>Simprocs\<close> |
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
975 |
|
65031
52e2c99f3711
use the cancellation simprocs directly
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
65029
diff
changeset
|
976 |
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
|
977 |
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
|
978 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
979 |
lemma mset_subseteq_add_iff1: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
980 |
"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
|
981 |
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
|
982 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
983 |
lemma mset_subseteq_add_iff2: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
984 |
"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
|
985 |
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
|
986 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
987 |
lemma mset_subset_add_iff1: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
988 |
"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
|
989 |
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
|
990 |
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
|
991 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
992 |
lemma mset_subset_add_iff2: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
993 |
"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
|
994 |
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
|
995 |
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
|
996 |
|
69605 | 997 |
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
|
998 |
|
65029
00731700e54f
cancellation simprocs generalising the multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
65027
diff
changeset
|
999 |
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
|
1000 |
by simp |
00731700e54f
cancellation simprocs generalising the multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
65027
diff
changeset
|
1001 |
|
00731700e54f
cancellation simprocs generalising the multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
65027
diff
changeset
|
1002 |
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
|
1003 |
|
00731700e54f
cancellation simprocs generalising the multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
65027
diff
changeset
|
1004 |
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
|
1005 |
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
|
1006 |
|
00731700e54f
cancellation simprocs generalising the multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
65027
diff
changeset
|
1007 |
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
|
1008 |
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
|
1009 |
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
|
1010 |
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
|
1011 |
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
|
1012 |
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
|
1013 |
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
|
1014 |
|
65027
2b8583507891
renaming multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
64911
diff
changeset
|
1015 |
simproc_setup mseteq_cancel |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1016 |
("(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
|
1017 |
"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
|
1018 |
"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
|
1019 |
"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
|
1020 |
\<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
|
1021 |
|
65027
2b8583507891
renaming multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
64911
diff
changeset
|
1022 |
simproc_setup msetsubset_cancel |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1023 |
("(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
|
1024 |
"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
|
1025 |
"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
|
1026 |
"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
|
1027 |
\<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
|
1028 |
|
65027
2b8583507891
renaming multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
64911
diff
changeset
|
1029 |
simproc_setup msetsubset_eq_cancel |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1030 |
("(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
|
1031 |
"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
|
1032 |
"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
|
1033 |
"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
|
1034 |
\<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
|
1035 |
|
65027
2b8583507891
renaming multiset simprocs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
64911
diff
changeset
|
1036 |
simproc_setup msetdiff_cancel |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1037 |
("((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
|
1038 |
"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
|
1039 |
"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
|
1040 |
"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
|
1041 |
\<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
|
1042 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1043 |
|
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1044 |
subsubsection \<open>Conditionally complete lattice\<close> |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1045 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1046 |
instantiation multiset :: (type) Inf |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1047 |
begin |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1048 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1049 |
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
|
1050 |
"\<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
|
1051 |
proof - |
73270 | 1052 |
fix A :: "('a \<Rightarrow> nat) set" |
1053 |
assume *: "\<And>f. f \<in> A \<Longrightarrow> finite {x. 0 < f x}" |
|
1054 |
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
|
1055 |
proof (cases "A = {}") |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1056 |
case False |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1057 |
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
|
1058 |
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
|
1059 |
by (auto intro: less_le_trans[OF _ cInf_lower]) |
73270 | 1060 |
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
|
1061 |
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
|
1062 |
with False show ?thesis by simp |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1063 |
qed simp_all |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1064 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1065 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1066 |
instance .. |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1067 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1068 |
end |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1069 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1070 |
lemma Inf_multiset_empty: "Inf {} = {#}" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1071 |
by transfer simp_all |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1072 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1073 |
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
|
1074 |
by transfer simp_all |
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 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1077 |
instantiation multiset :: (type) Sup |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1078 |
begin |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1079 |
|
63360 | 1080 |
definition Sup_multiset :: "'a multiset set \<Rightarrow> 'a multiset" where |
1081 |
"Sup_multiset A = (if A \<noteq> {} \<and> subset_mset.bdd_above A then |
|
1082 |
Abs_multiset (\<lambda>i. Sup ((\<lambda>X. count X i) ` A)) else {#})" |
|
1083 |
||
1084 |
lemma Sup_multiset_empty: "Sup {} = {#}" |
|
1085 |
by (simp add: Sup_multiset_def) |
|
1086 |
||
73451 | 1087 |
lemma Sup_multiset_unbounded: "\<not> subset_mset.bdd_above A \<Longrightarrow> Sup A = {#}" |
63360 | 1088 |
by (simp add: Sup_multiset_def) |
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1089 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1090 |
instance .. |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1091 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1092 |
end |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1093 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1094 |
lemma bdd_above_multiset_imp_bdd_above_count: |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1095 |
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
|
1096 |
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
|
1097 |
proof - |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1098 |
from assms obtain Y where Y: "\<forall>X\<in>A. X \<subseteq># Y" |
73451 | 1099 |
by (meson subset_mset.bdd_above.E) |
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1100 |
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
|
1101 |
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
|
1102 |
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
|
1103 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1104 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1105 |
lemma bdd_above_multiset_imp_finite_support: |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1106 |
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
|
1107 |
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
|
1108 |
proof - |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1109 |
from assms obtain Y where Y: "\<forall>X\<in>A. X \<subseteq># Y" |
73451 | 1110 |
by (meson subset_mset.bdd_above.E) |
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1111 |
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
|
1112 |
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
|
1113 |
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
|
1114 |
by safe (erule less_le_trans) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1115 |
moreover have "finite \<dots>" by simp |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1116 |
ultimately show ?thesis by (rule finite_subset) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1117 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1118 |
|
63360 | 1119 |
lemma Sup_multiset_in_multiset: |
73270 | 1120 |
\<open>finite {i. 0 < (SUP M\<in>A. count M i)}\<close> |
1121 |
if \<open>A \<noteq> {}\<close> \<open>subset_mset.bdd_above A\<close> |
|
1122 |
proof - |
|
63360 | 1123 |
have "{i. Sup ((\<lambda>X. count X i) ` A) > 0} \<subseteq> (\<Union>X\<in>A. {i. 0 < count X i})" |
1124 |
proof safe |
|
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69107
diff
changeset
|
1125 |
fix i assume pos: "(SUP X\<in>A. count X i) > 0" |
63360 | 1126 |
show "i \<in> (\<Union>X\<in>A. {i. 0 < count X i})" |
1127 |
proof (rule ccontr) |
|
1128 |
assume "i \<notin> (\<Union>X\<in>A. {i. 0 < count X i})" |
|
1129 |
hence "\<forall>X\<in>A. count X i \<le> 0" by (auto simp: count_eq_zero_iff) |
|
73270 | 1130 |
with that have "(SUP X\<in>A. count X i) \<le> 0" |
63360 | 1131 |
by (intro cSup_least bdd_above_multiset_imp_bdd_above_count) auto |
1132 |
with pos show False by simp |
|
1133 |
qed |
|
1134 |
qed |
|
73270 | 1135 |
moreover from that have "finite \<dots>" |
1136 |
by (rule bdd_above_multiset_imp_finite_support) |
|
1137 |
ultimately show "finite {i. Sup ((\<lambda>X. count X i) ` A) > 0}" |
|
1138 |
by (rule finite_subset) |
|
63360 | 1139 |
qed |
1140 |
||
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1141 |
lemma count_Sup_multiset_nonempty: |
73270 | 1142 |
\<open>count (Sup A) x = (SUP X\<in>A. count X x)\<close> |
1143 |
if \<open>A \<noteq> {}\<close> \<open>subset_mset.bdd_above A\<close> |
|
1144 |
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
|
1145 |
|
67398 | 1146 |
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
|
1147 |
proof |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1148 |
fix X :: "'a multiset" and A |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1149 |
assume "X \<in> A" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1150 |
show "Inf A \<subseteq># X" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1151 |
proof (rule mset_subset_eqI) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1152 |
fix x |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1153 |
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
|
1154 |
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
|
1155 |
by (simp add: count_Inf_multiset_nonempty) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1156 |
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
|
1157 |
by (intro cInf_lower) simp_all |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1158 |
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
|
1159 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1160 |
next |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1161 |
fix X :: "'a multiset" and A |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1162 |
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
|
1163 |
show "X \<subseteq># Inf A" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1164 |
proof (rule mset_subset_eqI) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1165 |
fix x |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69107
diff
changeset
|
1166 |
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
|
1167 |
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
|
1168 |
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
|
1169 |
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
|
1170 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1171 |
next |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1172 |
fix X :: "'a multiset" and A |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1173 |
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
|
1174 |
show "X \<subseteq># Sup A" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1175 |
proof (rule mset_subset_eqI) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1176 |
fix x |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1177 |
from X have "A \<noteq> {}" by auto |
69260
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
haftmann
parents:
69107
diff
changeset
|
1178 |
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
|
1179 |
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
|
1180 |
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
|
1181 |
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
|
1182 |
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
|
1183 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1184 |
next |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1185 |
fix X :: "'a multiset" and A |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1186 |
assume nonempty: "A \<noteq> {}" and ge: "\<And>Y. Y \<in> A \<Longrightarrow> Y \<subseteq># X" |
73451 | 1187 |
from ge have bdd: "subset_mset.bdd_above A" |
1188 |
by blast |
|
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1189 |
show "Sup A \<subseteq># X" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1190 |
proof (rule mset_subset_eqI) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1191 |
fix x |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1192 |
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
|
1193 |
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
|
1194 |
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
|
1195 |
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
|
1196 |
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
|
1197 |
qed |
64585
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
haftmann
parents:
64531
diff
changeset
|
1198 |
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
|
1199 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1200 |
lemma set_mset_Inf: |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1201 |
assumes "A \<noteq> {}" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1202 |
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
|
1203 |
proof safe |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1204 |
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
|
1205 |
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
|
1206 |
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
|
1207 |
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
|
1208 |
finally show "x \<in># X" by simp |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1209 |
next |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1210 |
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
|
1211 |
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
|
1212 |
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
|
1213 |
thus "x \<in># Inf A" by simp |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1214 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1215 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1216 |
lemma in_Inf_multiset_iff: |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1217 |
assumes "A \<noteq> {}" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1218 |
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
|
1219 |
proof - |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1220 |
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
|
1221 |
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
|
1222 |
finally show ?thesis . |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1223 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1224 |
|
63360 | 1225 |
lemma in_Inf_multisetD: "x \<in># Inf A \<Longrightarrow> X \<in> A \<Longrightarrow> x \<in># X" |
1226 |
by (subst (asm) in_Inf_multiset_iff) auto |
|
1227 |
||
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1228 |
lemma set_mset_Sup: |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1229 |
assumes "subset_mset.bdd_above A" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1230 |
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
|
1231 |
proof safe |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1232 |
fix x assume "x \<in># Sup A" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1233 |
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
|
1234 |
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
|
1235 |
proof (rule ccontr) |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1236 |
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
|
1237 |
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
|
1238 |
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
|
1239 |
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
|
1240 |
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
|
1241 |
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
|
1242 |
with \<open>x \<in># Sup A\<close> show False |
68406 | 1243 |
by (auto simp: subseteq_mset_def simp flip: count_greater_zero_iff |
1244 |
dest!: spec[of _ x]) |
|
63358
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1245 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1246 |
next |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1247 |
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
|
1248 |
hence "{#x#} \<subseteq># X" by auto |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1249 |
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
|
1250 |
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
|
1251 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1252 |
|
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1253 |
lemma in_Sup_multiset_iff: |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1254 |
assumes "subset_mset.bdd_above A" |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1255 |
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
|
1256 |
proof - |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1257 |
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
|
1258 |
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
|
1259 |
finally show ?thesis . |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1260 |
qed |
a500677d4cec
Conditionally complete lattice of multisets
Manuel Eberl <eberlm@in.tum.de>
parents:
63310
diff
changeset
|
1261 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1262 |
lemma in_Sup_multisetD: |
63360 | 1263 |
assumes "x \<in># Sup A" |
1264 |
shows "\<exists>X\<in>A. x \<in># X" |
|
1265 |
proof - |
|
1266 |
have "subset_mset.bdd_above A" |
|
1267 |
by (rule ccontr) (insert assms, simp_all add: Sup_multiset_unbounded) |
|
1268 |
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
|
1269 |
qed |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63524
diff
changeset
|
1270 |
|
67398 | 1271 |
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
|
1272 |
proof |
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63524
diff
changeset
|
1273 |
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
|
1274 |
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
|
1275 |
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
|
1276 |
qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> |
63360 | 1277 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1278 |
|
60500 | 1279 |
subsubsection \<open>Filter (with comprehension syntax)\<close> |
1280 |
||
1281 |
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
|
1282 |
|
59998
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
nipkow
parents:
59986
diff
changeset
|
1283 |
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
|
1284 |
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
|
1285 |
by (rule filter_preserves_multiset) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
1286 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1287 |
syntax (ASCII) |
63689 | 1288 |
"_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
|
1289 |
syntax |
63689 | 1290 |
"_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
|
1291 |
translations |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1292 |
"{#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
|
1293 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1294 |
lemma count_filter_mset [simp]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1295 |
"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
|
1296 |
by (simp add: filter_mset.rep_eq) |
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
nipkow
parents:
59986
diff
changeset
|
1297 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1298 |
lemma set_mset_filter [simp]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1299 |
"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
|
1300 |
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
|
1301 |
|
60606 | 1302 |
lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}" |
59998
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
nipkow
parents:
59986
diff
changeset
|
1303 |
by (rule multiset_eqI) simp |
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
nipkow
parents:
59986
diff
changeset
|
1304 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1305 |
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
|
1306 |
by (rule multiset_eqI) simp |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
1307 |
|
60606 | 1308 |
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
|
1309 |
by (rule multiset_eqI) simp |
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
haftmann
parents:
40968
diff
changeset
|
1310 |
|
60606 | 1311 |
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
|
1312 |
by (rule multiset_eqI) simp |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
1313 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
1314 |
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
|
1315 |
by (rule multiset_eqI) simp |
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
haftmann
parents:
40968
diff
changeset
|
1316 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
1317 |
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
|
1318 |
by (rule multiset_eqI) simp |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
1319 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1320 |
lemma filter_mset_add_mset [simp]: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1321 |
"filter_mset P (add_mset x A) = |
63795
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
1322 |
(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
|
1323 |
by (auto simp: multiset_eq_iff) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1324 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1325 |
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
|
1326 |
by (simp add: mset_subset_eqI) |
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59999
diff
changeset
|
1327 |
|
60606 | 1328 |
lemma multiset_filter_mono: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1329 |
assumes "A \<subseteq># B" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1330 |
shows "filter_mset f A \<subseteq># filter_mset f B" |
58035 | 1331 |
proof - |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
1332 |
from assms[unfolded mset_subset_eq_exists_conv] |
58035 | 1333 |
obtain C where B: "B = A + C" by auto |
1334 |
show ?thesis unfolding B by auto |
|
1335 |
qed |
|
1336 |
||
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1337 |
lemma filter_mset_eq_conv: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1338 |
"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
|
1339 |
proof |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1340 |
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
|
1341 |
next |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1342 |
assume ?Q |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1343 |
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
|
1344 |
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
|
1345 |
then have MN: "M - N = Q" by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1346 |
show ?P |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1347 |
proof (rule multiset_eqI) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1348 |
fix a |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1349 |
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
|
1350 |
by auto |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1351 |
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
|
1352 |
proof (cases "a \<in># M") |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1353 |
case True |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1354 |
with * show ?thesis |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1355 |
by (simp add: not_in_iff M) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1356 |
next |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1357 |
case False then have "count M a = 0" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1358 |
by (simp add: not_in_iff) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1359 |
with M show ?thesis by simp |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1360 |
qed |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1361 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1362 |
qed |
59813 | 1363 |
|
64077 | 1364 |
lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x \<in># M. Q x \<and> P x#}" |
1365 |
by (auto simp: multiset_eq_iff) |
|
1366 |
||
64418 | 1367 |
lemma |
1368 |
filter_mset_True[simp]: "{#y \<in># M. True#} = M" and |
|
1369 |
filter_mset_False[simp]: "{#y \<in># M. False#} = {#}" |
|
1370 |
by (auto simp: multiset_eq_iff) |
|
1371 |
||
75457 | 1372 |
lemma filter_mset_cong0: |
1373 |
assumes "\<And>x. x \<in># M \<Longrightarrow> f x \<longleftrightarrow> g x" |
|
1374 |
shows "filter_mset f M = filter_mset g M" |
|
1375 |
proof (rule subset_mset.antisym; unfold subseteq_mset_def; rule allI) |
|
1376 |
fix x |
|
1377 |
show "count (filter_mset f M) x \<le> count (filter_mset g M) x" |
|
1378 |
using assms by (cases "x \<in># M") (simp_all add: not_in_iff) |
|
1379 |
next |
|
1380 |
fix x |
|
1381 |
show "count (filter_mset g M) x \<le> count (filter_mset f M) x" |
|
1382 |
using assms by (cases "x \<in># M") (simp_all add: not_in_iff) |
|
1383 |
qed |
|
1384 |
||
1385 |
lemma filter_mset_cong: |
|
1386 |
assumes "M = M'" and "\<And>x. x \<in># M' \<Longrightarrow> f x \<longleftrightarrow> g x" |
|
1387 |
shows "filter_mset f M = filter_mset g M'" |
|
1388 |
unfolding \<open>M = M'\<close> |
|
1389 |
using assms by (auto intro: filter_mset_cong0) |
|
1390 |
||
79800 | 1391 |
lemma filter_eq_replicate_mset: "{#y \<in># D. y = x#} = replicate_mset (count D x) x" |
1392 |
by (induct D) (simp add: multiset_eqI) |
|
1393 |
||
59813 | 1394 |
|
60500 | 1395 |
subsubsection \<open>Size\<close> |
10249 | 1396 |
|
56656 | 1397 |
definition wcount where "wcount f M = (\<lambda>x. count M x * Suc (f x))" |
1398 |
||
1399 |
lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" |
|
1400 |
by (auto simp: wcount_def add_mult_distrib) |
|
1401 |
||
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1402 |
lemma wcount_add_mset: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1403 |
"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
|
1404 |
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
|
1405 |
|
56656 | 1406 |
definition size_multiset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a multiset \<Rightarrow> nat" where |
64267 | 1407 |
"size_multiset f M = sum (wcount f M) (set_mset M)" |
56656 | 1408 |
|
1409 |
lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] |
|
1410 |
||
60606 | 1411 |
instantiation multiset :: (type) size |
1412 |
begin |
|
1413 |
||
56656 | 1414 |
definition size_multiset where |
1415 |
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
|
1416 |
instance .. |
60606 | 1417 |
|
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
|
1418 |
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
|
1419 |
|
56656 | 1420 |
lemmas size_multiset_overloaded_eq = |
1421 |
size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] |
|
1422 |
||
1423 |
lemma size_multiset_empty [simp]: "size_multiset f {#} = 0" |
|
1424 |
by (simp add: size_multiset_def) |
|
1425 |
||
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
1426 |
lemma size_empty [simp]: "size {#} = 0" |
56656 | 1427 |
by (simp add: size_multiset_overloaded_def) |
1428 |
||
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1429 |
lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)" |
56656 | 1430 |
by (simp add: size_multiset_eq) |
10249 | 1431 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1432 |
lemma size_single: "size {#b#} = 1" |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1433 |
by (simp add: size_multiset_overloaded_def size_multiset_single) |
56656 | 1434 |
|
64267 | 1435 |
lemma sum_wcount_Int: |
1436 |
"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
|
1437 |
by (induct rule: finite_induct) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1438 |
(simp_all add: Int_insert_left wcount_def count_eq_zero_iff) |
56656 | 1439 |
|
1440 |
lemma size_multiset_union [simp]: |
|
1441 |
"size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" |
|
64267 | 1442 |
apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union) |
56656 | 1443 |
apply (subst Int_commute) |
64267 | 1444 |
apply (simp add: sum_wcount_Int) |
26178 | 1445 |
done |
10249 | 1446 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1447 |
lemma size_multiset_add_mset [simp]: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1448 |
"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
|
1449 |
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
|
1450 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1451 |
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
|
1452 |
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
|
1453 |
|
28708
a1a436f09ec6
explicit check for pattern discipline before code translation
haftmann
parents:
28562
diff
changeset
|
1454 |
lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" |
56656 | 1455 |
by (auto simp add: size_multiset_overloaded_def) |
1456 |
||
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1457 |
lemma size_multiset_eq_0_iff_empty [iff]: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1458 |
"size_multiset f M = 0 \<longleftrightarrow> M = {#}" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1459 |
by (auto simp add: size_multiset_eq count_eq_zero_iff) |
10249 | 1460 |
|
17161 | 1461 |
lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})" |
56656 | 1462 |
by (auto simp add: size_multiset_overloaded_def) |
26016 | 1463 |
|
1464 |
lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)" |
|
26178 | 1465 |
by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) |
10249 | 1466 |
|
60607 | 1467 |
lemma size_eq_Suc_imp_elem: "size M = Suc n \<Longrightarrow> \<exists>a. a \<in># M" |
56656 | 1468 |
apply (unfold size_multiset_overloaded_eq) |
64267 | 1469 |
apply (drule sum_SucD) |
26178 | 1470 |
apply auto |
1471 |
done |
|
10249 | 1472 |
|
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
|
1473 |
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
|
1474 |
assumes "size M = Suc n" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1475 |
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
|
1476 |
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
|
1477 |
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
|
1478 |
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
|
1479 |
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
|
1480 |
then show ?thesis by blast |
23611 | 1481 |
qed |
15869 | 1482 |
|
60606 | 1483 |
lemma size_mset_mono: |
1484 |
fixes A B :: "'a multiset" |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1485 |
assumes "A \<subseteq># B" |
60606 | 1486 |
shows "size A \<le> size B" |
59949 | 1487 |
proof - |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
1488 |
from assms[unfolded mset_subset_eq_exists_conv] |
59949 | 1489 |
obtain C where B: "B = A + C" by auto |
60606 | 1490 |
show ?thesis unfolding B by (induct C) auto |
59949 | 1491 |
qed |
1492 |
||
59998
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
nipkow
parents:
59986
diff
changeset
|
1493 |
lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \<le> size M" |
59949 | 1494 |
by (rule size_mset_mono[OF multiset_filter_subset]) |
1495 |
||
1496 |
lemma size_Diff_submset: |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1497 |
"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
|
1498 |
by (metis add_diff_cancel_left' size_union mset_subset_eq_exists_conv) |
26016 | 1499 |
|
79800 | 1500 |
lemma size_lt_imp_ex_count_lt: "size M < size N \<Longrightarrow> \<exists>x \<in># N. count M x < count N x" |
1501 |
by (metis count_eq_zero_iff leD not_le_imp_less not_less_zero size_mset_mono subseteq_mset_def) |
|
1502 |
||
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1503 |
|
60500 | 1504 |
subsection \<open>Induction and case splits\<close> |
10249 | 1505 |
|
18258 | 1506 |
theorem multiset_induct [case_names empty add, induct type: multiset]: |
48009 | 1507 |
assumes empty: "P {#}" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1508 |
assumes add: "\<And>x M. P M \<Longrightarrow> P (add_mset x M)" |
48009 | 1509 |
shows "P M" |
65545 | 1510 |
proof (induct "size M" arbitrary: M) |
48009 | 1511 |
case 0 thus "P M" by (simp add: empty) |
1512 |
next |
|
1513 |
case (Suc k) |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1514 |
obtain N x where "M = add_mset x N" |
60500 | 1515 |
using \<open>Suc k = size M\<close> [symmetric] |
48009 | 1516 |
using size_eq_Suc_imp_eq_union by fast |
1517 |
with Suc add show "P M" by simp |
|
10249 | 1518 |
qed |
1519 |
||
65545 | 1520 |
lemma multiset_induct_min[case_names empty add]: |
1521 |
fixes M :: "'a::linorder multiset" |
|
1522 |
assumes |
|
1523 |
empty: "P {#}" and |
|
1524 |
add: "\<And>x M. P M \<Longrightarrow> (\<forall>y \<in># M. y \<ge> x) \<Longrightarrow> P (add_mset x M)" |
|
1525 |
shows "P M" |
|
1526 |
proof (induct "size M" arbitrary: M) |
|
1527 |
case (Suc k) |
|
1528 |
note ih = this(1) and Sk_eq_sz_M = this(2) |
|
1529 |
||
66425 | 1530 |
let ?y = "Min_mset M" |
65545 | 1531 |
let ?N = "M - {#?y#}" |
1532 |
||
1533 |
have M: "M = add_mset ?y ?N" |
|
1534 |
by (metis Min_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero |
|
1535 |
set_mset_eq_empty_iff size_empty) |
|
1536 |
show ?case |
|
1537 |
by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, |
|
1538 |
meson Min_le finite_set_mset in_diffD) |
|
1539 |
qed (simp add: empty) |
|
1540 |
||
1541 |
lemma multiset_induct_max[case_names empty add]: |
|
1542 |
fixes M :: "'a::linorder multiset" |
|
1543 |
assumes |
|
1544 |
empty: "P {#}" and |
|
1545 |
add: "\<And>x M. P M \<Longrightarrow> (\<forall>y \<in># M. y \<le> x) \<Longrightarrow> P (add_mset x M)" |
|
1546 |
shows "P M" |
|
1547 |
proof (induct "size M" arbitrary: M) |
|
1548 |
case (Suc k) |
|
1549 |
note ih = this(1) and Sk_eq_sz_M = this(2) |
|
1550 |
||
66425 | 1551 |
let ?y = "Max_mset M" |
65545 | 1552 |
let ?N = "M - {#?y#}" |
1553 |
||
1554 |
have M: "M = add_mset ?y ?N" |
|
1555 |
by (metis Max_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero |
|
1556 |
set_mset_eq_empty_iff size_empty) |
|
1557 |
show ?case |
|
1558 |
by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, |
|
1559 |
meson Max_ge finite_set_mset in_diffD) |
|
1560 |
qed (simp add: empty) |
|
1561 |
||
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1562 |
lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = add_mset a A" |
26178 | 1563 |
by (induct M) auto |
25610 | 1564 |
|
55913 | 1565 |
lemma multiset_cases [cases type]: |
1566 |
obtains (empty) "M = {#}" |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1567 |
| (add) x N where "M = add_mset x N" |
63092 | 1568 |
by (induct M) simp_all |
25610 | 1569 |
|
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
|
1570 |
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
|
1571 |
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
|
1572 |
|
68992 | 1573 |
lemma union_filter_mset_complement[simp]: |
1574 |
"\<forall>x. P x = (\<not> Q x) \<Longrightarrow> filter_mset P M + filter_mset Q M = M" |
|
1575 |
by (subst multiset_eq_iff) auto |
|
1576 |
||
66494 | 1577 |
lemma multiset_partition: "M = {#x \<in># M. P x#} + {#x \<in># M. \<not> P x#}" |
68992 | 1578 |
by simp |
66494 | 1579 |
|
1580 |
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
|
1581 |
proof (induct A arbitrary: B) |
66494 | 1582 |
case empty |
1583 |
then show ?case |
|
1584 |
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
|
1585 |
next |
66494 | 1586 |
case (add x A) |
1587 |
have "add_mset x A \<subseteq># B" |
|
1588 |
by (meson add.prems subset_mset_def) |
|
1589 |
then show ?case |
|
1590 |
by (metis (no_types) add.prems add.right_neutral add_diff_cancel_left' leD nat_neq_iff |
|
1591 |
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
|
1592 |
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
|
1593 |
|
59949 | 1594 |
lemma size_1_singleton_mset: "size M = 1 \<Longrightarrow> \<exists>a. M = {#a#}" |
66494 | 1595 |
by (cases M) auto |
59949 | 1596 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1597 |
|
60500 | 1598 |
subsubsection \<open>Strong induction and subset induction for multisets\<close> |
1599 |
||
1600 |
text \<open>Well-foundedness of strict subset relation\<close> |
|
58098 | 1601 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
1602 |
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
|
1603 |
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
|
1604 |
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
|
1605 |
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
|
1606 |
|
76300 | 1607 |
lemma wfP_subset_mset[simp]: "wfP (\<subset>#)" |
1608 |
by (rule wf_subset_mset_rel[to_pred]) |
|
1609 |
||
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
|
1610 |
lemma full_multiset_induct [case_names less]: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1611 |
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
|
1612 |
shows "P B" |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
1613 |
apply (rule wf_subset_mset_rel [THEN wf_induct]) |
58098 | 1614 |
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
|
1615 |
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
|
1616 |
|
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 |
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
|
1618 |
assumes "F \<subseteq># A" |
60606 | 1619 |
and empty: "P {#}" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1620 |
and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (add_mset a F)" |
60606 | 1621 |
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
|
1622 |
proof - |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1623 |
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
|
1624 |
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
|
1625 |
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
|
1626 |
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
|
1627 |
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
|
1628 |
fix x F |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1629 |
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
|
1630 |
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
|
1631 |
proof (rule insert) |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
1632 |
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
|
1633 |
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
|
1634 |
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
|
1635 |
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
|
1636 |
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
|
1637 |
qed |
26145 | 1638 |
|
17161 | 1639 |
|
75467
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1640 |
subsection \<open>Least and greatest elements\<close> |
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1641 |
|
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1642 |
context begin |
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1643 |
|
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1644 |
qualified lemma |
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1645 |
assumes |
77699
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
desharna
parents:
77688
diff
changeset
|
1646 |
"M \<noteq> {#}" and |
76754
b5f4ae037fe2
used transp_on in assumptions of lemmas Multiset.bex_(least|greatest)_element
desharna
parents:
76749
diff
changeset
|
1647 |
"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
|
1648 |
"totalp_on (set_mset M) R" |
75467
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1649 |
shows |
77699
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
desharna
parents:
77688
diff
changeset
|
1650 |
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
|
1651 |
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
|
1652 |
using assms |
77699
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
desharna
parents:
77688
diff
changeset
|
1653 |
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
|
1654 |
|
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1655 |
end |
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1656 |
|
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
desharna
parents:
75459
diff
changeset
|
1657 |
|
60500 | 1658 |
subsection \<open>The fold combinator\<close> |
48023 | 1659 |
|
59998
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
nipkow
parents:
59986
diff
changeset
|
1660 |
definition fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b" |
48023 | 1661 |
where |
60495 | 1662 |
"fold_mset f s M = Finite_Set.fold (\<lambda>x. f x ^^ count M x) s (set_mset M)" |
48023 | 1663 |
|
60606 | 1664 |
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
|
1665 |
by (simp add: fold_mset_def) |
48023 | 1666 |
|
79800 | 1667 |
lemma fold_mset_single [simp]: "fold_mset f s {#x#} = f x s" |
1668 |
by (simp add: fold_mset_def) |
|
1669 |
||
48023 | 1670 |
context comp_fun_commute |
1671 |
begin |
|
1672 |
||
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1673 |
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
|
1674 |
proof - |
0cfc1651be25
simplified construction of fold combinator on multisets;
haftmann
parents:
49717
diff
changeset
|
1675 |
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
|
1676 |
by (fact comp_fun_commute_funpow) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1677 |
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
|
1678 |
by (fact comp_fun_commute_funpow) |
0cfc1651be25
simplified construction of fold combinator on multisets;
haftmann
parents:
49717
diff
changeset
|
1679 |
show ?thesis |
60495 | 1680 |
proof (cases "x \<in> set_mset M") |
49822
0cfc1651be25
simplified construction of fold combinator on multisets;
haftmann
parents:
49717
diff
changeset
|
1681 |
case False |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1682 |
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
|
1683 |
by (simp add: not_in_iff) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1684 |
from False have "Finite_Set.fold (\<lambda>y. f y ^^ count (add_mset x M) y) s (set_mset M) = |
60495 | 1685 |
Finite_Set.fold (\<lambda>y. f y ^^ count M y) s (set_mset M)" |
73832 | 1686 |
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
|
1687 |
with False * show ?thesis |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1688 |
by (simp add: fold_mset_def del: count_add_mset) |
48023 | 1689 |
next |
49822
0cfc1651be25
simplified construction of fold combinator on multisets;
haftmann
parents:
49717
diff
changeset
|
1690 |
case True |
63040 | 1691 |
define N where "N = set_mset M - {x}" |
60495 | 1692 |
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
|
1693 |
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
|
1694 |
Finite_Set.fold (\<lambda>y. f y ^^ count M y) s N" |
73832 | 1695 |
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
|
1696 |
with * show ?thesis by (simp add: fold_mset_def del: count_add_mset) simp |
48023 | 1697 |
qed |
1698 |
qed |
|
1699 |
||
60606 | 1700 |
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
|
1701 |
by (induct M) (simp_all add: fun_left_comm) |
48023 | 1702 |
|
60606 | 1703 |
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
|
1704 |
by (induct M) (simp_all add: fold_mset_fun_left_comm) |
48023 | 1705 |
|
1706 |
lemma fold_mset_fusion: |
|
1707 |
assumes "comp_fun_commute g" |
|
60606 | 1708 |
and *: "\<And>x y. h (g x y) = f x (h y)" |
1709 |
shows "h (fold_mset g w A) = fold_mset f (h w) A" |
|
48023 | 1710 |
proof - |
1711 |
interpret comp_fun_commute g by (fact assms) |
|
60606 | 1712 |
from * show ?thesis by (induct A) auto |
48023 | 1713 |
qed |
1714 |
||
1715 |
end |
|
1716 |
||
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1717 |
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
|
1718 |
proof - |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1719 |
interpret comp_fun_commute add_mset |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1720 |
by standard auto |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1721 |
show ?thesis |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1722 |
by (induction B) auto |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1723 |
qed |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1724 |
|
60500 | 1725 |
text \<open> |
48023 | 1726 |
A note on code generation: When defining some function containing a |
69593 | 1727 |
subterm \<^term>\<open>fold_mset F\<close>, code generation is not automatic. When |
61585 | 1728 |
interpreting locale \<open>left_commutative\<close> with \<open>F\<close>, the |
69593 | 1729 |
would be code thms for \<^const>\<open>fold_mset\<close> become thms like |
1730 |
\<^term>\<open>fold_mset F z {#} = z\<close> where \<open>F\<close> is not a pattern but |
|
48023 | 1731 |
contains defined symbols, i.e.\ is not a code thm. Hence a separate |
61585 | 1732 |
constant with its own code thms needs to be introduced for \<open>F\<close>. See the image operator below. |
60500 | 1733 |
\<close> |
1734 |
||
1735 |
||
1736 |
subsection \<open>Image\<close> |
|
48023 | 1737 |
|
1738 |
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
|
1739 |
"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
|
1740 |
|
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1741 |
lemma comp_fun_commute_mset_image: "comp_fun_commute (add_mset \<circ> f)" |
66494 | 1742 |
by unfold_locales (simp add: fun_eq_iff) |
48023 | 1743 |
|
1744 |
lemma image_mset_empty [simp]: "image_mset f {#} = {#}" |
|
49823 | 1745 |
by (simp add: image_mset_def) |
48023 | 1746 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1747 |
lemma image_mset_single: "image_mset f {#x#} = {#f x#}" |
66494 | 1748 |
by (simp add: comp_fun_commute.fold_mset_add_mset comp_fun_commute_mset_image image_mset_def) |
48023 | 1749 |
|
60606 | 1750 |
lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" |
49823 | 1751 |
proof - |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1752 |
interpret comp_fun_commute "add_mset \<circ> f" |
49823 | 1753 |
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
|
1754 |
show ?thesis by (induct N) (simp_all add: image_mset_def) |
49823 | 1755 |
qed |
1756 |
||
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1757 |
corollary image_mset_add_mset [simp]: |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1758 |
"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
|
1759 |
unfolding image_mset_union add_mset_add_single[of a M] by (simp add: image_mset_single) |
48023 | 1760 |
|
60606 | 1761 |
lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)" |
49823 | 1762 |
by (induct M) simp_all |
48040 | 1763 |
|
60606 | 1764 |
lemma size_image_mset [simp]: "size (image_mset f M) = size M" |
49823 | 1765 |
by (induct M) simp_all |
48023 | 1766 |
|
60606 | 1767 |
lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \<longleftrightarrow> M = {#}" |
49823 | 1768 |
by (cases M) auto |
48023 | 1769 |
|
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1770 |
lemma image_mset_If: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1771 |
"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
|
1772 |
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
|
1773 |
by (induction A) auto |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1774 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1775 |
lemma image_mset_Diff: |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1776 |
assumes "B \<subseteq># A" |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1777 |
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
|
1778 |
proof - |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1779 |
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
|
1780 |
by simp |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1781 |
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
|
1782 |
by (simp add: subset_mset.diff_add) |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1783 |
finally show ?thesis by simp |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1784 |
qed |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1785 |
|
73594 | 1786 |
lemma count_image_mset: |
1787 |
\<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
|
1788 |
proof (induction A) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1789 |
case empty |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1790 |
then show ?case by simp |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1791 |
next |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1792 |
case (add x A) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1793 |
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
|
1794 |
by simp |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1795 |
ultimately show ?case |
66494 | 1796 |
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
|
1797 |
qed |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
1798 |
|
73594 | 1799 |
lemma count_image_mset': |
1800 |
\<open>count (image_mset f X) y = (\<Sum>x | x \<in># X \<and> y = f x. count X x)\<close> |
|
1801 |
by (auto simp add: count_image_mset simp flip: singleton_conv2 simp add: Collect_conj_eq ac_simps) |
|
1802 |
||
63795
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
1803 |
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
|
1804 |
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
|
1805 |
|
65048 | 1806 |
lemma image_mset_subset_mono: "M \<subset># N \<Longrightarrow> image_mset f M \<subset># image_mset f N" |
1807 |
by (metis (no_types) Diff_eq_empty_iff_mset image_mset_Diff image_mset_is_empty_iff |
|
1808 |
image_mset_subseteq_mono subset_mset.less_le_not_le) |
|
1809 |
||
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
1810 |
syntax (ASCII) |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
1811 |
"_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" ("({#_/. _ :# _#})") |
48023 | 1812 |
syntax |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
1813 |
"_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" ("({#_/. _ \<in># _#})") |
59813 | 1814 |
translations |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
1815 |
"{#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
|
1816 |
|
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
1817 |
syntax (ASCII) |
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
1818 |
"_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" ("({#_/ | _ :# _./ _#})") |
48023 | 1819 |
syntax |
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
1820 |
"_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" ("({#_/ | _ \<in># _./ _#})") |
59813 | 1821 |
translations |
60606 | 1822 |
"{#e | x\<in>#M. P#}" \<rightharpoonup> "{#e. x \<in># {# x\<in>#M. P#}#}" |
59813 | 1823 |
|
60500 | 1824 |
text \<open> |
69593 | 1825 |
This allows to write not just filters like \<^term>\<open>{#x\<in>#M. x<c#}\<close> |
1826 |
but also images like \<^term>\<open>{#x+x. x\<in>#M #}\<close> and @{term [source] |
|
60607 | 1827 |
"{#x+x|x\<in>#M. x<c#}"}, where the latter is currently displayed as |
69593 | 1828 |
\<^term>\<open>{#x+x|x\<in>#M. x<c#}\<close>. |
60500 | 1829 |
\<close> |
48023 | 1830 |
|
60495 | 1831 |
lemma in_image_mset: "y \<in># {#f x. x \<in># M#} \<longleftrightarrow> y \<in> f ` set_mset M" |
66494 | 1832 |
by simp |
59813 | 1833 |
|
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
|
1834 |
functor image_mset: image_mset |
48023 | 1835 |
proof - |
1836 |
fix f g show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)" |
|
1837 |
proof |
|
1838 |
fix A |
|
1839 |
show "(image_mset f \<circ> image_mset g) A = image_mset (f \<circ> g) A" |
|
1840 |
by (induct A) simp_all |
|
1841 |
qed |
|
1842 |
show "image_mset id = id" |
|
1843 |
proof |
|
1844 |
fix A |
|
1845 |
show "image_mset id A = id A" |
|
1846 |
by (induct A) simp_all |
|
1847 |
qed |
|
1848 |
qed |
|
1849 |
||
59813 | 1850 |
declare |
1851 |
image_mset.id [simp] |
|
1852 |
image_mset.identity [simp] |
|
1853 |
||
1854 |
lemma image_mset_id[simp]: "image_mset id x = x" |
|
1855 |
unfolding id_def by auto |
|
1856 |
||
1857 |
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#}" |
|
1858 |
by (induct M) auto |
|
1859 |
||
1860 |
lemma image_mset_cong_pair: |
|
1861 |
"(\<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#}" |
|
1862 |
by (metis image_mset_cong split_cong) |
|
49717 | 1863 |
|
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64587
diff
changeset
|
1864 |
lemma image_mset_const_eq: |
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64587
diff
changeset
|
1865 |
"{#c. a \<in># M#} = replicate_mset (size M) c" |
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64587
diff
changeset
|
1866 |
by (induct M) simp_all |
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64587
diff
changeset
|
1867 |
|
75459 | 1868 |
lemma image_mset_filter_mset_swap: |
1869 |
"image_mset f (filter_mset (\<lambda>x. P (f x)) M) = filter_mset P (image_mset f M)" |
|
1870 |
by (induction M rule: multiset_induct) simp_all |
|
1871 |
||
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
|
1872 |
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
|
1873 |
"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
|
1874 |
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
|
1875 |
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
|
1876 |
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
|
1877 |
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
|
1878 |
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
|
1879 |
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
|
1880 |
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
|
1881 |
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
|
1882 |
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
|
1883 |
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
|
1884 |
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
|
1885 |
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
|
1886 |
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
|
1887 |
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
|
1888 |
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
|
1889 |
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
|
1890 |
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
|
1891 |
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
|
1892 |
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
|
1893 |
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
|
1894 |
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
|
1895 |
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
|
1896 |
|
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 |
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
|
1898 |
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
|
1899 |
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
|
1900 |
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
|
1901 |
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
|
1902 |
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
|
1903 |
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
|
1904 |
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
|
1905 |
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
|
1906 |
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
|
1907 |
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
|
1908 |
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
|
1909 |
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
|
1910 |
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
|
1911 |
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
|
1912 |
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
|
1913 |
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
|
1914 |
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
|
1915 |
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
|
1916 |
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
|
1917 |
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
|
1918 |
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
|
1919 |
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
|
1920 |
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
|
1921 |
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
|
1922 |
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
|
1923 |
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
|
1924 |
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
|
1925 |
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
|
1926 |
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
|
1927 |
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
|
1928 |
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
|
1929 |
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
|
1930 |
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
|
1931 |
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
|
1932 |
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
|
1933 |
|
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
|
1934 |
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
|
1935 |
"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
|
1936 |
\<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
|
1937 |
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
|
1938 |
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
|
1939 |
|
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
|
1940 |
|
60500 | 1941 |
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
|
1942 |
|
60515 | 1943 |
primrec mset :: "'a list \<Rightarrow> 'a multiset" where |
1944 |
"mset [] = {#}" | |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
1945 |
"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
|
1946 |
|
37107 | 1947 |
lemma in_multiset_in_set: |
60515 | 1948 |
"x \<in># mset xs \<longleftrightarrow> x \<in> set xs" |
37107 | 1949 |
by (induct xs) simp_all |
1950 |
||
60515 | 1951 |
lemma count_mset: |
1952 |
"count (mset xs) x = length (filter (\<lambda>y. x = y) xs)" |
|
37107 | 1953 |
by (induct xs) simp_all |
1954 |
||
60515 | 1955 |
lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])" |
59813 | 1956 |
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
|
1957 |
|
60515 | 1958 |
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
|
1959 |
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
|
1960 |
|
66276
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
1961 |
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
|
1962 |
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
|
1963 |
|
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
1964 |
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
|
1965 |
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
|
1966 |
|
64077 | 1967 |
lemma mset_single_iff[iff]: "mset xs = {#x#} \<longleftrightarrow> xs = [x]" |
1968 |
by (cases xs) auto |
|
1969 |
||
1970 |
lemma mset_single_iff_right[iff]: "{#x#} = mset xs \<longleftrightarrow> xs = [x]" |
|
1971 |
by (cases xs) auto |
|
1972 |
||
64076 | 1973 |
lemma set_mset_mset[simp]: "set_mset (mset xs) = set xs" |
1974 |
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
|
1975 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
1976 |
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
|
1977 |
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
|
1978 |
|
60515 | 1979 |
lemma size_mset [simp]: "size (mset xs) = length xs" |
48012 | 1980 |
by (induct xs) simp_all |
1981 |
||
60606 | 1982 |
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
|
1983 |
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
|
1984 |
|
68988 | 1985 |
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
|
1986 |
by (induct xs) simp_all |
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
1987 |
|
60515 | 1988 |
lemma mset_rev [simp]: |
1989 |
"mset (rev xs) = mset xs" |
|
40950 | 1990 |
by (induct xs) simp_all |
1991 |
||
60515 | 1992 |
lemma surj_mset: "surj mset" |
76359 | 1993 |
unfolding surj_def |
1994 |
proof (rule allI) |
|
1995 |
fix M |
|
1996 |
show "\<exists>xs. M = mset xs" |
|
1997 |
by (induction M) (auto intro: exI[of _ "_ # _"]) |
|
1998 |
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
|
1999 |
|
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
|
2000 |
lemma distinct_count_atmost_1: |
60606 | 2001 |
"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
|
2002 |
proof (induct x) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2003 |
case Nil then show ?case by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2004 |
next |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2005 |
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
|
2006 |
proof |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2007 |
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
|
2008 |
next |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2009 |
assume ?rhs then have "x \<notin> set xs" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2010 |
by (simp split: if_splits) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2011 |
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
|
2012 |
(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
|
2013 |
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
|
2014 |
ultimately show ?lhs using Cons by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2015 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2016 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2017 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2018 |
lemma mset_eq_setD: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2019 |
assumes "mset xs = mset ys" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2020 |
shows "set xs = set ys" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2021 |
proof - |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2022 |
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
|
2023 |
by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2024 |
then show ?thesis by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2025 |
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
|
2026 |
|
60515 | 2027 |
lemma set_eq_iff_mset_eq_distinct: |
73301 | 2028 |
\<open>distinct x \<Longrightarrow> distinct y \<Longrightarrow> set x = set y \<longleftrightarrow> mset x = mset y\<close> |
2029 |
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
|
2030 |
|
60515 | 2031 |
lemma set_eq_iff_mset_remdups_eq: |
73301 | 2032 |
\<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
|
2033 |
apply (rule iffI) |
60515 | 2034 |
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
|
2035 |
apply (drule distinct_remdups [THEN distinct_remdups |
60515 | 2036 |
[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
|
2037 |
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
|
2038 |
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
|
2039 |
|
73301 | 2040 |
lemma mset_eq_imp_distinct_iff: |
2041 |
\<open>distinct xs \<longleftrightarrow> distinct ys\<close> if \<open>mset xs = mset ys\<close> |
|
2042 |
using that by (auto simp add: distinct_count_atmost_1 dest: mset_eq_setD) |
|
2043 |
||
60607 | 2044 |
lemma nth_mem_mset: "i < length ls \<Longrightarrow> (ls ! i) \<in># mset ls" |
60678 | 2045 |
proof (induct ls arbitrary: i) |
2046 |
case Nil |
|
2047 |
then show ?case by simp |
|
2048 |
next |
|
2049 |
case Cons |
|
2050 |
then show ?case by (cases i) auto |
|
2051 |
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
|
2052 |
|
60606 | 2053 |
lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}" |
60678 | 2054 |
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
|
2055 |
|
60515 | 2056 |
lemma mset_eq_length: |
2057 |
assumes "mset xs = mset ys" |
|
37107 | 2058 |
shows "length xs = length ys" |
60515 | 2059 |
using assms by (metis size_mset) |
2060 |
||
2061 |
lemma mset_eq_length_filter: |
|
2062 |
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
|
2063 |
shows "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) ys)" |
60515 | 2064 |
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
|
2065 |
|
45989
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
haftmann
parents:
45866
diff
changeset
|
2066 |
lemma fold_multiset_equiv: |
73706 | 2067 |
\<open>List.fold f xs = List.fold f ys\<close> |
2068 |
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> |
|
2069 |
and \<open>mset xs = mset ys\<close> |
|
2070 |
using f \<open>mset xs = mset ys\<close> [symmetric] proof (induction xs arbitrary: ys) |
|
60678 | 2071 |
case Nil |
2072 |
then show ?case by simp |
|
45989
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
haftmann
parents:
45866
diff
changeset
|
2073 |
next |
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
haftmann
parents:
45866
diff
changeset
|
2074 |
case (Cons x xs) |
73706 | 2075 |
then have *: \<open>set ys = set (x # xs)\<close> |
60678 | 2076 |
by (blast dest: mset_eq_setD) |
73706 | 2077 |
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
|
2078 |
by (rule Cons.prems(1)) (simp_all add: *) |
73706 | 2079 |
moreover from * have \<open>x \<in> set ys\<close> |
2080 |
by simp |
|
2081 |
ultimately have \<open>List.fold f ys = List.fold f (remove1 x ys) \<circ> f x\<close> |
|
2082 |
by (fact fold_remove1_split) |
|
2083 |
moreover from Cons.prems have \<open>List.fold f xs = List.fold f (remove1 x ys)\<close> |
|
2084 |
by (auto intro: Cons.IH) |
|
2085 |
ultimately show ?case |
|
60678 | 2086 |
by simp |
73706 | 2087 |
qed |
2088 |
||
2089 |
lemma fold_permuted_eq: |
|
2090 |
\<open>List.fold (\<odot>) xs z = List.fold (\<odot>) ys z\<close> |
|
2091 |
if \<open>mset xs = mset ys\<close> |
|
2092 |
and \<open>P z\<close> and P: \<open>\<And>x z. x \<in> set xs \<Longrightarrow> P z \<Longrightarrow> P (x \<odot> z)\<close> |
|
2093 |
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> |
|
2094 |
for f (infixl \<open>\<odot>\<close> 70) |
|
2095 |
using \<open>P z\<close> P f \<open>mset xs = mset ys\<close> [symmetric] proof (induction xs arbitrary: ys z) |
|
2096 |
case Nil |
|
2097 |
then show ?case by simp |
|
2098 |
next |
|
2099 |
case (Cons x xs) |
|
2100 |
then have *: \<open>set ys = set (x # xs)\<close> |
|
2101 |
by (blast dest: mset_eq_setD) |
|
2102 |
have \<open>P z\<close> |
|
2103 |
by (fact Cons.prems(1)) |
|
2104 |
moreover have \<open>\<And>x z. x \<in> set ys \<Longrightarrow> P z \<Longrightarrow> P (x \<odot> z)\<close> |
|
2105 |
by (rule Cons.prems(2)) (simp_all add: *) |
|
2106 |
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> |
|
2107 |
by (rule Cons.prems(3)) (simp_all add: *) |
|
2108 |
moreover from * have \<open>x \<in> set ys\<close> |
|
2109 |
by simp |
|
2110 |
ultimately have \<open>fold (\<odot>) ys z = fold (\<odot>) (remove1 x ys) (x \<odot> z)\<close> |
|
2111 |
by (induction ys arbitrary: z) auto |
|
2112 |
moreover from Cons.prems have \<open>fold (\<odot>) xs (x \<odot> z) = fold (\<odot>) (remove1 x ys) (x \<odot> z)\<close> |
|
2113 |
by (auto intro: Cons.IH) |
|
2114 |
ultimately show ?case |
|
2115 |
by simp |
|
45989
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
haftmann
parents:
45866
diff
changeset
|
2116 |
qed |
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
haftmann
parents:
45866
diff
changeset
|
2117 |
|
69107 | 2118 |
lemma mset_shuffles: "zs \<in> shuffles xs ys \<Longrightarrow> mset zs = mset xs + mset ys" |
2119 |
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
|
2120 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2121 |
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
|
2122 |
by (induct xs) simp_all |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2123 |
|
63524
4ec755485732
adding mset_map to the simp rules
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63410
diff
changeset
|
2124 |
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
|
2125 |
by (induct xs) simp_all |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
2126 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2127 |
global_interpretation mset_set: folding add_mset "{#}" |
73832 | 2128 |
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
|
2129 |
by standard (simp add: fun_eq_iff) |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2130 |
|
66276
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2131 |
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
|
2132 |
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
|
2133 |
|
60513 | 2134 |
lemma count_mset_set [simp]: |
2135 |
"finite A \<Longrightarrow> x \<in> A \<Longrightarrow> count (mset_set A) x = 1" (is "PROP ?P") |
|
2136 |
"\<not> finite A \<Longrightarrow> count (mset_set A) x = 0" (is "PROP ?Q") |
|
2137 |
"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
|
2138 |
proof - |
60606 | 2139 |
have *: "count (mset_set A) x = 0" if "x \<notin> A" for A |
2140 |
proof (cases "finite A") |
|
2141 |
case False then show ?thesis by simp |
|
2142 |
next |
|
2143 |
case True from True \<open>x \<notin> A\<close> show ?thesis by (induct A) auto |
|
2144 |
qed |
|
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
2145 |
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
|
2146 |
by (auto elim!: Set.set_insert) |
69593 | 2147 |
qed \<comment> \<open>TODO: maybe define \<^const>\<open>mset_set\<close> also in terms of \<^const>\<open>Abs_multiset\<close>\<close> |
60513 | 2148 |
|
2149 |
lemma elem_mset_set[simp, intro]: "finite A \<Longrightarrow> x \<in># mset_set A \<longleftrightarrow> x \<in> A" |
|
59813 | 2150 |
by (induct A rule: finite_induct) simp_all |
2151 |
||
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2152 |
lemma mset_set_Union: |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2153 |
"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
|
2154 |
by (induction A rule: finite_induct) auto |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2155 |
|
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2156 |
lemma filter_mset_mset_set [simp]: |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2157 |
"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
|
2158 |
proof (induction A rule: finite_induct) |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2159 |
case (insert x A) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2160 |
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
|
2161 |
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
|
2162 |
by simp |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2163 |
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
|
2164 |
by (rule insert.IH) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2165 |
also from insert.hyps |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2166 |
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
|
2167 |
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
|
2168 |
by (intro mset_set_Union [symmetric]) simp_all |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2169 |
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
|
2170 |
finally show ?case . |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2171 |
qed simp_all |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2172 |
|
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2173 |
lemma mset_set_Diff: |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2174 |
assumes "finite A" "B \<subseteq> A" |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2175 |
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
|
2176 |
proof - |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2177 |
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
|
2178 |
by (intro mset_set_Union) (auto dest: finite_subset) |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2179 |
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
|
2180 |
finally show ?thesis by simp |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2181 |
qed |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2182 |
|
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2183 |
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
|
2184 |
by (induction xs) simp_all |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2185 |
|
66276
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2186 |
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
|
2187 |
by auto |
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2188 |
|
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2189 |
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
|
2190 |
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
|
2191 |
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
|
2192 |
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
|
2193 |
fix x :: 'a |
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2194 |
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
|
2195 |
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
|
2196 |
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
|
2197 |
qed |
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2198 |
|
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2199 |
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
|
2200 |
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
|
2201 |
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
|
2202 |
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
|
2203 |
qed |
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
eberlm <eberlm@in.tum.de>
parents:
65547
diff
changeset
|
2204 |
|
73466 | 2205 |
lemma mset_set_upto_eq_mset_upto: |
2206 |
\<open>mset_set {..<n} = mset [0..<n]\<close> |
|
2207 |
by (induction n) (auto simp: ac_simps lessThan_Suc) |
|
2208 |
||
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2209 |
context linorder |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2210 |
begin |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2211 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2212 |
definition sorted_list_of_multiset :: "'a multiset \<Rightarrow> 'a list" |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2213 |
where |
59998
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
nipkow
parents:
59986
diff
changeset
|
2214 |
"sorted_list_of_multiset M = fold_mset insort [] M" |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2215 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2216 |
lemma sorted_list_of_multiset_empty [simp]: |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2217 |
"sorted_list_of_multiset {#} = []" |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2218 |
by (simp add: sorted_list_of_multiset_def) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2219 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2220 |
lemma sorted_list_of_multiset_singleton [simp]: |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2221 |
"sorted_list_of_multiset {#x#} = [x]" |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2222 |
proof - |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2223 |
interpret comp_fun_commute insort by (fact comp_fun_commute_insort) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2224 |
show ?thesis by (simp add: sorted_list_of_multiset_def) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2225 |
qed |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2226 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2227 |
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
|
2228 |
"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
|
2229 |
proof - |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2230 |
interpret comp_fun_commute insort by (fact comp_fun_commute_insort) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2231 |
show ?thesis by (simp add: sorted_list_of_multiset_def) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2232 |
qed |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2233 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2234 |
end |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2235 |
|
66494 | 2236 |
lemma mset_sorted_list_of_multiset[simp]: "mset (sorted_list_of_multiset M) = M" |
2237 |
by (induct M) simp_all |
|
2238 |
||
2239 |
lemma sorted_list_of_multiset_mset[simp]: "sorted_list_of_multiset (mset xs) = sort xs" |
|
2240 |
by (induct xs) simp_all |
|
2241 |
||
2242 |
lemma finite_set_mset_mset_set[simp]: "finite A \<Longrightarrow> set_mset (mset_set A) = A" |
|
2243 |
by auto |
|
60513 | 2244 |
|
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2245 |
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
|
2246 |
using finite_set_mset_mset_set by fastforce |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2247 |
|
66494 | 2248 |
lemma infinite_set_mset_mset_set: "\<not> finite A \<Longrightarrow> set_mset (mset_set A) = {}" |
2249 |
by simp |
|
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2250 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2251 |
lemma set_sorted_list_of_multiset [simp]: |
60495 | 2252 |
"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
|
2253 |
by (induct M) (simp_all add: set_insort_key) |
60513 | 2254 |
|
2255 |
lemma sorted_list_of_mset_set [simp]: |
|
2256 |
"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
|
2257 |
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
|
2258 |
|
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2259 |
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
|
2260 |
by (induction n) (simp_all add: atLeastLessThanSuc) |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2261 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2262 |
lemma image_mset_map_of: |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2263 |
"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
|
2264 |
proof (induction xs) |
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2265 |
case (Cons x xs) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2266 |
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
|
2267 |
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
|
2268 |
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
|
2269 |
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
|
2270 |
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
|
2271 |
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
|
2272 |
finally show ?case by simp |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2273 |
qed simp_all |
63099
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
eberlm
parents:
63092
diff
changeset
|
2274 |
|
66494 | 2275 |
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
|
2276 |
assumes "finite A" "finite B" |
66494 | 2277 |
shows "mset_set A \<subseteq># mset_set B \<longleftrightarrow> A \<subseteq> B" |
2278 |
using assms set_mset_mono subset_imp_msubset_mset_set by fastforce |
|
2279 |
||
2280 |
lemma mset_set_eq_iff[simp]: |
|
2281 |
assumes "finite A" "finite B" |
|
2282 |
shows "mset_set A = mset_set B \<longleftrightarrow> A = B" |
|
2283 |
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
|
2284 |
|
69895
6b03a8cf092d
more formal contributors (with the help of the history);
wenzelm
parents:
69605
diff
changeset
|
2285 |
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
|
2286 |
assumes "inj_on f A" |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2287 |
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
|
2288 |
proof cases |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2289 |
assume "finite A" |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2290 |
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
|
2291 |
by (induct A) auto |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2292 |
next |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2293 |
assume "infinite A" |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2294 |
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
|
2295 |
using finite_imageD by blast |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2296 |
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
|
2297 |
qed |
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
eberlm <eberlm@in.tum.de>
parents:
63908
diff
changeset
|
2298 |
|
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2299 |
|
79800 | 2300 |
subsection \<open>More properties of the replicate, repeat, and image operations\<close> |
60804 | 2301 |
|
2302 |
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
|
2303 |
unfolding replicate_mset_def by (induct n) auto |
60804 | 2304 |
|
2305 |
lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})" |
|
2306 |
by (auto split: if_splits) |
|
2307 |
||
2308 |
lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" |
|
2309 |
by (induct n, simp_all) |
|
2310 |
||
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
2311 |
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
|
2312 |
by (auto simp add: mset_subset_eqI) (metis count_replicate_mset subseteq_mset_def) |
60804 | 2313 |
|
66494 | 2314 |
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
|
2315 |
by (induct xs) auto |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2316 |
|
66494 | 2317 |
lemma replicate_mset_eq_empty_iff [simp]: "replicate_mset n a = {#} \<longleftrightarrow> n = 0" |
62366 | 2318 |
by (induct n) simp_all |
2319 |
||
2320 |
lemma replicate_mset_eq_iff: |
|
66494 | 2321 |
"replicate_mset m a = replicate_mset n b \<longleftrightarrow> m = 0 \<and> n = 0 \<or> m = n \<and> a = b" |
62366 | 2322 |
by (auto simp add: multiset_eq_iff) |
2323 |
||
63908
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
2324 |
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
|
2325 |
by (auto simp: multiset_eq_iff) |
0dd6731060d7
delete looping simp rule
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63831
diff
changeset
|
2326 |
|
63908
ca41b6670904
support replicate_mset in multiset simproc
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63882
diff
changeset
|
2327 |
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
|
2328 |
by (auto simp: multiset_eq_iff) |
0dd6731060d7
delete looping simp rule
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63831
diff
changeset
|
2329 |
|
64077 | 2330 |
lemma repeat_mset_eq_empty_iff: "repeat_mset n A = {#} \<longleftrightarrow> n = 0 \<or> A = {#}" |
2331 |
by (cases n) auto |
|
2332 |
||
63924 | 2333 |
lemma image_replicate_mset [simp]: |
2334 |
"image_mset f (replicate_mset n a) = replicate_mset n (f a)" |
|
2335 |
by (induct n) simp_all |
|
2336 |
||
67051 | 2337 |
lemma replicate_mset_msubseteq_iff: |
2338 |
"replicate_mset m a \<subseteq># replicate_mset n b \<longleftrightarrow> m = 0 \<or> a = b \<and> m \<le> n" |
|
2339 |
by (cases m) |
|
68406 | 2340 |
(auto simp: insert_subset_eq_iff simp flip: count_le_replicate_mset_subset_eq) |
67051 | 2341 |
|
2342 |
lemma msubseteq_replicate_msetE: |
|
2343 |
assumes "A \<subseteq># replicate_mset n a" |
|
2344 |
obtains m where "m \<le> n" and "A = replicate_mset m a" |
|
2345 |
proof (cases "n = 0") |
|
2346 |
case True |
|
2347 |
with assms that show thesis |
|
2348 |
by simp |
|
2349 |
next |
|
2350 |
case False |
|
2351 |
from assms have "set_mset A \<subseteq> set_mset (replicate_mset n a)" |
|
2352 |
by (rule set_mset_mono) |
|
2353 |
with False have "set_mset A \<subseteq> {a}" |
|
2354 |
by simp |
|
2355 |
then have "\<exists>m. A = replicate_mset m a" |
|
2356 |
proof (induction A) |
|
2357 |
case empty |
|
2358 |
then show ?case |
|
2359 |
by simp |
|
2360 |
next |
|
2361 |
case (add b A) |
|
2362 |
then obtain m where "A = replicate_mset m a" |
|
2363 |
by auto |
|
2364 |
with add.prems show ?case |
|
2365 |
by (auto intro: exI [of _ "Suc m"]) |
|
2366 |
qed |
|
2367 |
then obtain m where A: "A = replicate_mset m a" .. |
|
2368 |
with assms have "m \<le> n" |
|
2369 |
by (auto simp add: replicate_mset_msubseteq_iff) |
|
2370 |
then show thesis using A .. |
|
2371 |
qed |
|
2372 |
||
79800 | 2373 |
lemma count_image_mset_lt_imp_lt_raw: |
2374 |
assumes |
|
2375 |
"finite A" and |
|
2376 |
"A = set_mset M \<union> set_mset N" and |
|
2377 |
"count (image_mset f M) b < count (image_mset f N) b" |
|
2378 |
shows "\<exists>x. f x = b \<and> count M x < count N x" |
|
2379 |
using assms |
|
2380 |
proof (induct A arbitrary: M N b rule: finite_induct) |
|
2381 |
case (insert x F) |
|
2382 |
note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and |
|
2383 |
cnt_b = this(5) |
|
2384 |
||
2385 |
let ?Ma = "{#y \<in># M. y \<noteq> x#}" |
|
2386 |
let ?Mb = "{#y \<in># M. y = x#}" |
|
2387 |
let ?Na = "{#y \<in># N. y \<noteq> x#}" |
|
2388 |
let ?Nb = "{#y \<in># N. y = x#}" |
|
2389 |
||
2390 |
have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" |
|
2391 |
using multiset_partition by blast+ |
|
2392 |
||
2393 |
have f_eq_ma_na: "F = set_mset ?Ma \<union> set_mset ?Na" |
|
2394 |
using x_f_eq_m_n x_ni_f by auto |
|
2395 |
||
2396 |
show ?case |
|
2397 |
proof (cases "count (image_mset f ?Ma) b < count (image_mset f ?Na) b") |
|
2398 |
case cnt_ba: True |
|
2399 |
obtain xa where "f xa = b" and "count ?Ma xa < count ?Na xa" |
|
2400 |
using ih[OF f_eq_ma_na cnt_ba] by blast |
|
2401 |
thus ?thesis |
|
2402 |
by (metis count_filter_mset not_less0) |
|
2403 |
next |
|
2404 |
case cnt_ba: False |
|
2405 |
have fx_eq_b: "f x = b" |
|
2406 |
using cnt_b cnt_ba |
|
2407 |
by (subst (asm) m_part, subst (asm) n_part, |
|
2408 |
auto simp: filter_eq_replicate_mset split: if_splits) |
|
2409 |
moreover have "count M x < count N x" |
|
2410 |
using cnt_b cnt_ba |
|
2411 |
by (subst (asm) m_part, subst (asm) n_part, |
|
2412 |
auto simp: filter_eq_replicate_mset split: if_splits) |
|
2413 |
ultimately show ?thesis |
|
2414 |
by blast |
|
2415 |
qed |
|
2416 |
qed auto |
|
2417 |
||
2418 |
lemma count_image_mset_lt_imp_lt: |
|
2419 |
assumes cnt_b: "count (image_mset f M) b < count (image_mset f N) b" |
|
2420 |
shows "\<exists>x. f x = b \<and> count M x < count N x" |
|
2421 |
by (rule count_image_mset_lt_imp_lt_raw[of "set_mset M \<union> set_mset N", OF _ refl cnt_b]) auto |
|
2422 |
||
2423 |
lemma count_image_mset_le_imp_lt_raw: |
|
2424 |
assumes |
|
2425 |
"finite A" and |
|
2426 |
"A = set_mset M \<union> set_mset N" and |
|
2427 |
"count (image_mset f M) (f a) + count N a < count (image_mset f N) (f a) + count M a" |
|
2428 |
shows "\<exists>b. f b = f a \<and> count M b < count N b" |
|
2429 |
using assms |
|
2430 |
proof (induct A arbitrary: M N rule: finite_induct) |
|
2431 |
case (insert x F) |
|
2432 |
note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and |
|
2433 |
cnt_lt = this(5) |
|
2434 |
||
2435 |
let ?Ma = "{#y \<in># M. y \<noteq> x#}" |
|
2436 |
let ?Mb = "{#y \<in># M. y = x#}" |
|
2437 |
let ?Na = "{#y \<in># N. y \<noteq> x#}" |
|
2438 |
let ?Nb = "{#y \<in># N. y = x#}" |
|
2439 |
||
2440 |
have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" |
|
2441 |
using multiset_partition by blast+ |
|
2442 |
||
2443 |
have f_eq_ma_na: "F = set_mset ?Ma \<union> set_mset ?Na" |
|
2444 |
using x_f_eq_m_n x_ni_f by auto |
|
2445 |
||
2446 |
show ?case |
|
2447 |
proof (cases "f x = f a") |
|
2448 |
case fx_ne_fa: False |
|
2449 |
||
2450 |
have cnt_fma_fa: "count (image_mset f ?Ma) (f a) = count (image_mset f M) (f a)" |
|
2451 |
using fx_ne_fa by (subst (2) m_part) (auto simp: filter_eq_replicate_mset) |
|
2452 |
have cnt_fna_fa: "count (image_mset f ?Na) (f a) = count (image_mset f N) (f a)" |
|
2453 |
using fx_ne_fa by (subst (2) n_part) (auto simp: filter_eq_replicate_mset) |
|
2454 |
have cnt_ma_a: "count ?Ma a = count M a" |
|
2455 |
using fx_ne_fa by (subst (2) m_part) (auto simp: filter_eq_replicate_mset) |
|
2456 |
have cnt_na_a: "count ?Na a = count N a" |
|
2457 |
using fx_ne_fa by (subst (2) n_part) (auto simp: filter_eq_replicate_mset) |
|
2458 |
||
2459 |
obtain b where fb_eq_fa: "f b = f a" and cnt_b: "count ?Ma b < count ?Na b" |
|
2460 |
using ih[OF f_eq_ma_na] cnt_lt unfolding cnt_fma_fa cnt_fna_fa cnt_ma_a cnt_na_a by blast |
|
2461 |
have fx_ne_fb: "f x \<noteq> f b" |
|
2462 |
using fb_eq_fa fx_ne_fa by simp |
|
2463 |
||
2464 |
have cnt_ma_b: "count ?Ma b = count M b" |
|
2465 |
using fx_ne_fb by (subst (2) m_part) auto |
|
2466 |
have cnt_na_b: "count ?Na b = count N b" |
|
2467 |
using fx_ne_fb by (subst (2) n_part) auto |
|
2468 |
||
2469 |
show ?thesis |
|
2470 |
using fb_eq_fa cnt_b unfolding cnt_ma_b cnt_na_b by blast |
|
2471 |
next |
|
2472 |
case fx_eq_fa: True |
|
2473 |
show ?thesis |
|
2474 |
proof (cases "x = a") |
|
2475 |
case x_eq_a: True |
|
2476 |
have "count (image_mset f ?Ma) (f a) + count ?Na a |
|
2477 |
< count (image_mset f ?Na) (f a) + count ?Ma a" |
|
2478 |
using cnt_lt x_eq_a by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, |
|
2479 |
auto simp: filter_eq_replicate_mset) |
|
2480 |
thus ?thesis |
|
2481 |
using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) |
|
2482 |
next |
|
2483 |
case x_ne_a: False |
|
2484 |
show ?thesis |
|
2485 |
proof (cases "count M x < count N x") |
|
2486 |
case True |
|
2487 |
thus ?thesis |
|
2488 |
using fx_eq_fa by blast |
|
2489 |
next |
|
2490 |
case False |
|
2491 |
hence cnt_x: "count M x \<ge> count N x" |
|
2492 |
by fastforce |
|
2493 |
have "count M x + count (image_mset f ?Ma) (f a) + count ?Na a |
|
2494 |
< count N x + count (image_mset f ?Na) (f a) + count ?Ma a" |
|
2495 |
using cnt_lt x_ne_a fx_eq_fa by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, |
|
2496 |
auto simp: filter_eq_replicate_mset) |
|
2497 |
hence "count (image_mset f ?Ma) (f a) + count ?Na a |
|
2498 |
< count (image_mset f ?Na) (f a) + count ?Ma a" |
|
2499 |
using cnt_x by linarith |
|
2500 |
thus ?thesis |
|
2501 |
using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) |
|
2502 |
qed |
|
2503 |
qed |
|
2504 |
qed |
|
2505 |
qed auto |
|
2506 |
||
2507 |
lemma count_image_mset_le_imp_lt: |
|
2508 |
assumes |
|
2509 |
"count (image_mset f M) (f a) \<le> count (image_mset f N) (f a)" and |
|
2510 |
"count M a > count N a" |
|
2511 |
shows "\<exists>b. f b = f a \<and> count M b < count N b" |
|
2512 |
using assms by (auto intro: count_image_mset_le_imp_lt_raw[of "set_mset M \<union> set_mset N"]) |
|
2513 |
||
2514 |
lemma size_filter_unsat_elem: |
|
2515 |
assumes "x \<in># M" and "\<not> P x" |
|
2516 |
shows "size {#x \<in># M. P x#} < size M" |
|
2517 |
proof - |
|
2518 |
have "size (filter_mset P M) \<noteq> size M" |
|
2519 |
using assms by (metis add.right_neutral add_diff_cancel_left' count_filter_mset mem_Collect_eq |
|
2520 |
multiset_partition nonempty_has_size set_mset_def size_union) |
|
2521 |
then show ?thesis |
|
2522 |
by (meson leD nat_neq_iff size_filter_mset_lesseq) |
|
2523 |
qed |
|
2524 |
||
2525 |
lemma size_filter_ne_elem: "x \<in># M \<Longrightarrow> size {#y \<in># M. y \<noteq> x#} < size M" |
|
2526 |
by (simp add: size_filter_unsat_elem[of x M "\<lambda>y. y \<noteq> x"]) |
|
2527 |
||
2528 |
lemma size_eq_ex_count_lt: |
|
2529 |
assumes |
|
2530 |
sz_m_eq_n: "size M = size N" and |
|
2531 |
m_eq_n: "M \<noteq> N" |
|
2532 |
shows "\<exists>x. count M x < count N x" |
|
2533 |
proof - |
|
2534 |
obtain x where "count M x \<noteq> count N x" |
|
2535 |
using m_eq_n by (meson multiset_eqI) |
|
2536 |
moreover |
|
2537 |
{ |
|
2538 |
assume "count M x < count N x" |
|
2539 |
hence ?thesis |
|
2540 |
by blast |
|
2541 |
} |
|
2542 |
moreover |
|
2543 |
{ |
|
2544 |
assume cnt_x: "count M x > count N x" |
|
2545 |
||
2546 |
have "size {#y \<in># M. y = x#} + size {#y \<in># M. y \<noteq> x#} = |
|
2547 |
size {#y \<in># N. y = x#} + size {#y \<in># N. y \<noteq> x#}" |
|
2548 |
using sz_m_eq_n multiset_partition by (metis size_union) |
|
2549 |
hence sz_m_minus_x: "size {#y \<in># M. y \<noteq> x#} < size {#y \<in># N. y \<noteq> x#}" |
|
2550 |
using cnt_x by (simp add: filter_eq_replicate_mset) |
|
2551 |
then obtain y where "count {#y \<in># M. y \<noteq> x#} y < count {#y \<in># N. y \<noteq> x#} y" |
|
2552 |
using size_lt_imp_ex_count_lt[OF sz_m_minus_x] by blast |
|
2553 |
hence "count M y < count N y" |
|
2554 |
by (metis count_filter_mset less_asym) |
|
2555 |
hence ?thesis |
|
2556 |
by blast |
|
2557 |
} |
|
2558 |
ultimately show ?thesis |
|
2559 |
by fastforce |
|
2560 |
qed |
|
2561 |
||
60804 | 2562 |
|
60500 | 2563 |
subsection \<open>Big operators\<close> |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2564 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2565 |
locale comm_monoid_mset = comm_monoid |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2566 |
begin |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2567 |
|
64075 | 2568 |
interpretation comp_fun_commute f |
2569 |
by standard (simp add: fun_eq_iff left_commute) |
|
2570 |
||
2571 |
interpretation comp?: comp_fun_commute "f \<circ> g" |
|
2572 |
by (fact comp_comp_fun_commute) |
|
2573 |
||
2574 |
context |
|
2575 |
begin |
|
2576 |
||
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2577 |
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
|
2578 |
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
|
2579 |
|
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
63195
diff
changeset
|
2580 |
lemma empty [simp]: "F {#} = \<^bold>1" |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2581 |
by (simp add: eq_fold) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2582 |
|
60678 | 2583 |
lemma singleton [simp]: "F {#x#} = x" |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2584 |
proof - |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2585 |
interpret comp_fun_commute |
60678 | 2586 |
by standard (simp add: fun_eq_iff left_commute) |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2587 |
show ?thesis by (simp add: eq_fold) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2588 |
qed |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2589 |
|
63290
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
haftmann
parents:
63195
diff
changeset
|
2590 |
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
|
2591 |
proof - |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2592 |
interpret comp_fun_commute f |
60678 | 2593 |
by standard (simp add: fun_eq_iff left_commute) |
2594 |
show ?thesis |
|
2595 |
by (induct N) (simp_all add: left_commute eq_fold) |
|
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2596 |
qed |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2597 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2598 |
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
|
2599 |
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
|
2600 |
|
64075 | 2601 |
lemma insert [simp]: |
2602 |
shows "F (image_mset g (add_mset x A)) = g x \<^bold>* F (image_mset g A)" |
|
2603 |
by (simp add: eq_fold) |
|
2604 |
||
2605 |
lemma remove: |
|
2606 |
assumes "x \<in># A" |
|
2607 |
shows "F A = x \<^bold>* F (A - {#x#})" |
|
2608 |
using multi_member_split[OF assms] by auto |
|
2609 |
||
2610 |
lemma neutral: |
|
2611 |
"\<forall>x\<in>#A. x = \<^bold>1 \<Longrightarrow> F A = \<^bold>1" |
|
2612 |
by (induct A) simp_all |
|
2613 |
||
2614 |
lemma neutral_const [simp]: |
|
2615 |
"F (image_mset (\<lambda>_. \<^bold>1) A) = \<^bold>1" |
|
2616 |
by (simp add: neutral) |
|
2617 |
||
2618 |
private lemma F_image_mset_product: |
|
2619 |
"F {#g x j \<^bold>* F {#g i j. i \<in># A#}. j \<in># B#} = |
|
2620 |
F (image_mset (g x) B) \<^bold>* F {#F {#g i j. i \<in># A#}. j \<in># B#}" |
|
2621 |
by (induction B) (simp_all add: left_commute semigroup.assoc semigroup_axioms) |
|
2622 |
||
68938 | 2623 |
lemma swap: |
64075 | 2624 |
"F (image_mset (\<lambda>i. F (image_mset (g i) B)) A) = |
2625 |
F (image_mset (\<lambda>j. F (image_mset (\<lambda>i. g i j) A)) B)" |
|
2626 |
apply (induction A, simp) |
|
2627 |
apply (induction B, auto simp add: F_image_mset_product ac_simps) |
|
2628 |
done |
|
2629 |
||
2630 |
lemma distrib: "F (image_mset (\<lambda>x. g x \<^bold>* h x) A) = F (image_mset g A) \<^bold>* F (image_mset h A)" |
|
2631 |
by (induction A) (auto simp: ac_simps) |
|
2632 |
||
2633 |
lemma union_disjoint: |
|
2634 |
"A \<inter># B = {#} \<Longrightarrow> F (image_mset g (A \<union># B)) = F (image_mset g A) \<^bold>* F (image_mset g B)" |
|
2635 |
by (induction A) (auto simp: ac_simps) |
|
2636 |
||
2637 |
end |
|
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2638 |
end |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2639 |
|
67398 | 2640 |
lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute ((+) :: 'a multiset \<Rightarrow> _ \<Rightarrow> _)" |
60678 | 2641 |
by standard (simp add: add_ac comp_def) |
59813 | 2642 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2643 |
declare comp_fun_commute.fold_mset_add_mset[OF comp_fun_commute_plus_mset, simp] |
59813 | 2644 |
|
67398 | 2645 |
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 | 2646 |
by (induct NN) auto |
2647 |
||
54868 | 2648 |
context comm_monoid_add |
2649 |
begin |
|
2650 |
||
63830 | 2651 |
sublocale sum_mset: comm_monoid_mset plus 0 |
2652 |
defines sum_mset = sum_mset.F .. |
|
2653 |
||
64267 | 2654 |
lemma sum_unfold_sum_mset: |
2655 |
"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
|
2656 |
by (cases "finite A") (induct A rule: finite_induct, simp_all) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2657 |
|
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2658 |
end |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2659 |
|
73047
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
72607
diff
changeset
|
2660 |
notation sum_mset ("\<Sum>\<^sub>#") |
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
72607
diff
changeset
|
2661 |
|
62366 | 2662 |
syntax (ASCII) |
63830 | 2663 |
"_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3SUM _:#_. _)" [0, 51, 10] 10) |
62366 | 2664 |
syntax |
63830 | 2665 |
"_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10) |
62366 | 2666 |
translations |
63830 | 2667 |
"\<Sum>i \<in># A. b" \<rightleftharpoons> "CONST sum_mset (CONST image_mset (\<lambda>i. b) A)" |
59949 | 2668 |
|
66938 | 2669 |
context comm_monoid_add |
2670 |
begin |
|
2671 |
||
2672 |
lemma sum_mset_sum_list: |
|
2673 |
"sum_mset (mset xs) = sum_list xs" |
|
2674 |
by (induction xs) auto |
|
2675 |
||
2676 |
end |
|
2677 |
||
2678 |
context canonically_ordered_monoid_add |
|
2679 |
begin |
|
2680 |
||
2681 |
lemma sum_mset_0_iff [simp]: |
|
2682 |
"sum_mset M = 0 \<longleftrightarrow> (\<forall>x \<in> set_mset M. x = 0)" |
|
2683 |
by (induction M) auto |
|
2684 |
||
2685 |
end |
|
2686 |
||
2687 |
context ordered_comm_monoid_add |
|
2688 |
begin |
|
2689 |
||
2690 |
lemma sum_mset_mono: |
|
2691 |
"sum_mset (image_mset f K) \<le> sum_mset (image_mset g K)" |
|
2692 |
if "\<And>i. i \<in># K \<Longrightarrow> f i \<le> g i" |
|
2693 |
using that by (induction K) (simp_all add: add_mono) |
|
2694 |
||
2695 |
end |
|
2696 |
||
73470 | 2697 |
context cancel_comm_monoid_add |
66938 | 2698 |
begin |
2699 |
||
2700 |
lemma sum_mset_diff: |
|
2701 |
"sum_mset (M - N) = sum_mset M - sum_mset N" if "N \<subseteq># M" for M N :: "'a multiset" |
|
2702 |
using that by (auto simp add: subset_mset.le_iff_add) |
|
2703 |
||
2704 |
end |
|
2705 |
||
2706 |
context semiring_0 |
|
2707 |
begin |
|
2708 |
||
63860 | 2709 |
lemma sum_mset_distrib_left: |
66938 | 2710 |
"c * (\<Sum>x \<in># M. f x) = (\<Sum>x \<in># M. c * f(x))" |
2711 |
by (induction M) (simp_all add: algebra_simps) |
|
63860 | 2712 |
|
64075 | 2713 |
lemma sum_mset_distrib_right: |
66938 | 2714 |
"(\<Sum>x \<in># M. f x) * c = (\<Sum>x \<in># M. f x * c)" |
2715 |
by (induction M) (simp_all add: algebra_simps) |
|
2716 |
||
2717 |
end |
|
2718 |
||
2719 |
lemma sum_mset_product: |
|
2720 |
fixes f :: "'a::{comm_monoid_add,times} \<Rightarrow> 'b::semiring_0" |
|
2721 |
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 | 2722 |
by (subst sum_mset.swap) (simp add: sum_mset_distrib_left sum_mset_distrib_right) |
66938 | 2723 |
|
2724 |
context semiring_1 |
|
2725 |
begin |
|
2726 |
||
2727 |
lemma sum_mset_replicate_mset [simp]: |
|
2728 |
"sum_mset (replicate_mset n a) = of_nat n * a" |
|
2729 |
by (induction n) (simp_all add: algebra_simps) |
|
2730 |
||
2731 |
lemma sum_mset_delta: |
|
2732 |
"sum_mset (image_mset (\<lambda>x. if x = y then c else 0) A) = c * of_nat (count A y)" |
|
2733 |
by (induction A) (simp_all add: algebra_simps) |
|
2734 |
||
2735 |
lemma sum_mset_delta': |
|
2736 |
"sum_mset (image_mset (\<lambda>x. if y = x then c else 0) A) = c * of_nat (count A y)" |
|
2737 |
by (induction A) (simp_all add: algebra_simps) |
|
2738 |
||
2739 |
end |
|
2740 |
||
2741 |
lemma of_nat_sum_mset [simp]: |
|
2742 |
"of_nat (sum_mset A) = sum_mset (image_mset of_nat A)" |
|
2743 |
by (induction A) auto |
|
2744 |
||
2745 |
lemma size_eq_sum_mset: |
|
2746 |
"size M = (\<Sum>a\<in>#M. 1)" |
|
2747 |
using image_mset_const_eq [of "1::nat" M] by simp |
|
2748 |
||
2749 |
lemma size_mset_set [simp]: |
|
2750 |
"size (mset_set A) = card A" |
|
2751 |
by (simp only: size_eq_sum_mset card_eq_sum sum_unfold_sum_mset) |
|
64075 | 2752 |
|
2753 |
lemma sum_mset_constant [simp]: |
|
2754 |
fixes y :: "'b::semiring_1" |
|
2755 |
shows \<open>(\<Sum>x\<in>#A. y) = of_nat (size A) * y\<close> |
|
2756 |
by (induction A) (auto simp: algebra_simps) |
|
2757 |
||
73047
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
72607
diff
changeset
|
2758 |
lemma set_mset_Union_mset[simp]: "set_mset (\<Sum>\<^sub># MM) = (\<Union>M \<in> set_mset MM. set_mset M)" |
59813 | 2759 |
by (induct MM) auto |
2760 |
||
73047
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
72607
diff
changeset
|
2761 |
lemma in_Union_mset_iff[iff]: "x \<in># \<Sum>\<^sub># MM \<longleftrightarrow> (\<exists>M. M \<in># MM \<and> x \<in># M)" |
59813 | 2762 |
by (induct MM) auto |
2763 |
||
64267 | 2764 |
lemma count_sum: |
2765 |
"count (sum f A) x = sum (\<lambda>a. count (f a) x) A" |
|
62366 | 2766 |
by (induct A rule: infinite_finite_induct) simp_all |
2767 |
||
64267 | 2768 |
lemma sum_eq_empty_iff: |
62366 | 2769 |
assumes "finite A" |
64267 | 2770 |
shows "sum f A = {#} \<longleftrightarrow> (\<forall>a\<in>A. f a = {#})" |
62366 | 2771 |
using assms by induct simp_all |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2772 |
|
73047
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
72607
diff
changeset
|
2773 |
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
|
2774 |
by (induction M) auto |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
2775 |
|
73047
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
72607
diff
changeset
|
2776 |
lemma Union_image_single_mset[simp]: "\<Sum>\<^sub># (image_mset (\<lambda>x. {#x#}) m) = m" |
67656 | 2777 |
by(induction m) auto |
2778 |
||
54868 | 2779 |
context comm_monoid_mult |
2780 |
begin |
|
2781 |
||
63830 | 2782 |
sublocale prod_mset: comm_monoid_mset times 1 |
2783 |
defines prod_mset = prod_mset.F .. |
|
2784 |
||
2785 |
lemma prod_mset_empty: |
|
2786 |
"prod_mset {#} = 1" |
|
2787 |
by (fact prod_mset.empty) |
|
2788 |
||
2789 |
lemma prod_mset_singleton: |
|
2790 |
"prod_mset {#x#} = x" |
|
2791 |
by (fact prod_mset.singleton) |
|
2792 |
||
2793 |
lemma prod_mset_Un: |
|
2794 |
"prod_mset (A + B) = prod_mset A * prod_mset B" |
|
2795 |
by (fact prod_mset.union) |
|
2796 |
||
66938 | 2797 |
lemma prod_mset_prod_list: |
2798 |
"prod_mset (mset xs) = prod_list xs" |
|
2799 |
by (induct xs) auto |
|
2800 |
||
63830 | 2801 |
lemma prod_mset_replicate_mset [simp]: |
2802 |
"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
|
2803 |
by (induct n) simp_all |
60804 | 2804 |
|
64272 | 2805 |
lemma prod_unfold_prod_mset: |
2806 |
"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
|
2807 |
by (cases "finite A") (induct A rule: finite_induct, simp_all) |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2808 |
|
63830 | 2809 |
lemma prod_mset_multiplicity: |
64272 | 2810 |
"prod_mset M = prod (\<lambda>x. x ^ count M x) (set_mset M)" |
2811 |
by (simp add: fold_mset_def prod.eq_fold prod_mset.eq_fold funpow_times_power comp_def) |
|
63830 | 2812 |
|
2813 |
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
|
2814 |
by (induction A) simp_all |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63524
diff
changeset
|
2815 |
|
63830 | 2816 |
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
|
2817 |
by (induction A) simp_all |
63534
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
eberlm <eberlm@in.tum.de>
parents:
63524
diff
changeset
|
2818 |
|
66938 | 2819 |
lemma prod_mset_subset_imp_dvd: |
2820 |
assumes "A \<subseteq># B" |
|
2821 |
shows "prod_mset A dvd prod_mset B" |
|
2822 |
proof - |
|
2823 |
from assms have "B = (B - A) + A" by (simp add: subset_mset.diff_add) |
|
2824 |
also have "prod_mset \<dots> = prod_mset (B - A) * prod_mset A" by simp |
|
2825 |
also have "prod_mset A dvd \<dots>" by simp |
|
2826 |
finally show ?thesis . |
|
2827 |
qed |
|
2828 |
||
2829 |
lemma dvd_prod_mset: |
|
2830 |
assumes "x \<in># A" |
|
2831 |
shows "x dvd prod_mset A" |
|
2832 |
using assms prod_mset_subset_imp_dvd [of "{#x#}" A] by simp |
|
2833 |
||
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2834 |
end |
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2835 |
|
73052
c03a148110cc
HOL-Library.Multiset: new notation for prod_mset, consistent with sum_mset
Manuel Eberl <eberlm@in.tum.de>
parents:
73047
diff
changeset
|
2836 |
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
|
2837 |
|
61955
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
wenzelm
parents:
61890
diff
changeset
|
2838 |
syntax (ASCII) |
63830 | 2839 |
"_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
|
2840 |
syntax |
63830 | 2841 |
"_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
|
2842 |
translations |
63830 | 2843 |
"\<Prod>i \<in># A. b" \<rightleftharpoons> "CONST prod_mset (CONST image_mset (\<lambda>i. b) A)" |
2844 |
||
64591
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64587
diff
changeset
|
2845 |
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
|
2846 |
by (simp add: image_mset_const_eq) |
240a39af9ec4
restructured matter on polynomials and normalized fractions
haftmann
parents:
64587
diff
changeset
|
2847 |
|
63830 | 2848 |
lemma (in semidom) prod_mset_zero_iff [iff]: |
2849 |
"prod_mset A = 0 \<longleftrightarrow> 0 \<in># A" |
|
62366 | 2850 |
by (induct A) auto |
2851 |
||
63830 | 2852 |
lemma (in semidom_divide) prod_mset_diff: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2853 |
assumes "B \<subseteq># A" and "0 \<notin># B" |
63830 | 2854 |
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
|
2855 |
proof - |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2856 |
from assms obtain C where "A = B + C" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2857 |
by (metis subset_mset.add_diff_inverse) |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2858 |
with assms show ?thesis by simp |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2859 |
qed |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2860 |
|
63830 | 2861 |
lemma (in semidom_divide) prod_mset_minus: |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2862 |
assumes "a \<in># A" and "a \<noteq> 0" |
63830 | 2863 |
shows "prod_mset (A - {#a#}) = prod_mset A div a" |
2864 |
using assms prod_mset_diff [of "{#a#}" A] by auto |
|
2865 |
||
71398
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2866 |
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
|
2867 |
"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
|
2868 |
proof (induction A) |
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2869 |
case (add x A) |
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2870 |
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
|
2871 |
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
|
2872 |
by simp |
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2873 |
also note add.IH |
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2874 |
finally show ?case by simp |
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2875 |
qed auto |
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2876 |
|
63924 | 2877 |
lemma (in algebraic_semidom) is_unit_prod_mset_iff: |
2878 |
"is_unit (prod_mset A) \<longleftrightarrow> (\<forall>x \<in># A. is_unit x)" |
|
2879 |
by (induct A) (auto simp: is_unit_mult_iff) |
|
2880 |
||
71398
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2881 |
lemma (in normalization_semidom_multiplicative) normalize_prod_mset: |
63924 | 2882 |
"normalize (prod_mset A) = prod_mset (image_mset normalize A)" |
2883 |
by (induct A) (simp_all add: normalize_mult) |
|
2884 |
||
71398
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
Manuel Eberl <eberlm@in.tum.de>
parents:
69895
diff
changeset
|
2885 |
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
|
2886 |
assumes "\<And>a. a \<in># A \<Longrightarrow> normalize a = a" |
63830 | 2887 |
shows "normalize (prod_mset A) = prod_mset A" |
63924 | 2888 |
proof - |
2889 |
from assms have "image_mset normalize A = A" |
|
2890 |
by (induct A) simp_all |
|
2891 |
then show ?thesis by (simp add: normalize_prod_mset) |
|
2892 |
qed |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2893 |
|
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2894 |
|
73301 | 2895 |
subsection \<open>Multiset as order-ignorant lists\<close> |
51548
757fa47af981
centralized various multiset operations in theory multiset;
haftmann
parents:
51161
diff
changeset
|
2896 |
|
39533
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2897 |
context linorder |
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2898 |
begin |
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2899 |
|
60515 | 2900 |
lemma mset_insort [simp]: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
2901 |
"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
|
2902 |
by (induct xs) simp_all |
58425 | 2903 |
|
60515 | 2904 |
lemma mset_sort [simp]: |
2905 |
"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
|
2906 |
by (induct xs) simp_all |
37107 | 2907 |
|
60500 | 2908 |
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
|
2909 |
This lemma shows which properties suffice to show that a function |
61585 | 2910 |
\<open>f\<close> with \<open>f xs = ys\<close> behaves like sort. |
60500 | 2911 |
\<close> |
37074 | 2912 |
|
39533
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2913 |
lemma properties_for_sort_key: |
60515 | 2914 |
assumes "mset ys = mset xs" |
60606 | 2915 |
and "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>x. f k = f x) ys = filter (\<lambda>x. f k = f x) xs" |
2916 |
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
|
2917 |
shows "sort_key f xs = ys" |
60606 | 2918 |
using assms |
46921 | 2919 |
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
|
2920 |
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
|
2921 |
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
|
2922 |
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
|
2923 |
from Cons.prems(2) have |
40305
41833242cc42
tuned lemma proposition of properties_for_sort_key
haftmann
parents:
40303
diff
changeset
|
2924 |
"\<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
|
2925 |
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
|
2926 |
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
|
2927 |
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
|
2928 |
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
|
2929 |
by auto |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2930 |
then have "x \<in> set ys" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
2931 |
by simp |
39533
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2932 |
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
|
2933 |
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
|
2934 |
|
39533
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2935 |
lemma properties_for_sort: |
60515 | 2936 |
assumes multiset: "mset ys = mset xs" |
60606 | 2937 |
and "sorted ys" |
39533
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2938 |
shows "sort xs = ys" |
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2939 |
proof (rule properties_for_sort_key) |
60515 | 2940 |
from multiset show "mset ys = mset xs" . |
60500 | 2941 |
from \<open>sorted ys\<close> show "sorted (map (\<lambda>x. x) ys)" by simp |
60678 | 2942 |
from multiset have "length (filter (\<lambda>y. k = y) ys) = length (filter (\<lambda>x. k = x) xs)" for k |
60515 | 2943 |
by (rule mset_eq_length_filter) |
60678 | 2944 |
then have "replicate (length (filter (\<lambda>y. k = y) ys)) k = |
2945 |
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
|
2946 |
by simp |
60678 | 2947 |
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
|
2948 |
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
|
2949 |
qed |
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
2950 |
|
61031
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2951 |
lemma sort_key_inj_key_eq: |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2952 |
assumes mset_equal: "mset xs = mset ys" |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2953 |
and "inj_on f (set xs)" |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2954 |
and "sorted (map f ys)" |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2955 |
shows "sort_key f xs = ys" |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2956 |
proof (rule properties_for_sort_key) |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2957 |
from mset_equal |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2958 |
show "mset ys = mset xs" by simp |
61188 | 2959 |
from \<open>sorted (map f ys)\<close> |
61031
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2960 |
show "sorted (map f ys)" . |
68386 | 2961 |
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
|
2962 |
proof - |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2963 |
from mset_equal |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2964 |
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
|
2965 |
with that have "insert k (set ys) = set ys" by auto |
61188 | 2966 |
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
|
2967 |
by (simp add: set_equal) |
68386 | 2968 |
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
|
2969 |
by (auto intro!: inj_on_filter_key_eq) |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2970 |
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
|
2971 |
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
|
2972 |
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
|
2973 |
using mset_equal by simp |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2974 |
also have "\<dots> = filter (HOL.eq k) xs" |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2975 |
by (simp add: replicate_count_mset_eq_filter_eq) |
68386 | 2976 |
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
|
2977 |
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
|
2978 |
finally show ?thesis . |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2979 |
qed |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2980 |
qed |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2981 |
|
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2982 |
lemma sort_key_eq_sort_key: |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2983 |
assumes "mset xs = mset ys" |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2984 |
and "inj_on f (set xs)" |
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
haftmann
parents:
60804
diff
changeset
|
2985 |
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
|
2986 |
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
|
2987 |
|
40303
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
2988 |
lemma sort_key_by_quicksort: |
68386 | 2989 |
"sort_key f xs = sort_key f [x\<leftarrow>xs. f x < f (xs ! (length xs div 2))] |
2990 |
@ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))] |
|
2991 |
@ 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
|
2992 |
proof (rule properties_for_sort_key) |
60515 | 2993 |
show "mset ?rhs = mset ?lhs" |
69442 | 2994 |
by (rule multiset_eqI) auto |
40303
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
2995 |
show "sorted (map f ?rhs)" |
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
2996 |
by (auto simp add: sorted_append intro: sorted_map_same) |
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
2997 |
next |
40305
41833242cc42
tuned lemma proposition of properties_for_sort_key
haftmann
parents:
40303
diff
changeset
|
2998 |
fix l |
41833242cc42
tuned lemma proposition of properties_for_sort_key
haftmann
parents:
40303
diff
changeset
|
2999 |
assume "l \<in> set ?rhs" |
40346 | 3000 |
let ?pivot = "f (xs ! (length xs div 2))" |
3001 |
have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto |
|
68386 | 3002 |
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
|
3003 |
unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) |
68386 | 3004 |
with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp |
40346 | 3005 |
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 | 3006 |
then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] = |
3007 |
[x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp |
|
67398 | 3008 |
note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"] |
68386 | 3009 |
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
|
3010 |
proof (cases "f l" ?pivot rule: linorder_cases) |
46730 | 3011 |
case less |
3012 |
then have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto |
|
3013 |
with less show ?thesis |
|
40346 | 3014 |
by (simp add: filter_sort [symmetric] ** ***) |
40305
41833242cc42
tuned lemma proposition of properties_for_sort_key
haftmann
parents:
40303
diff
changeset
|
3015 |
next |
40306 | 3016 |
case equal then show ?thesis |
40346 | 3017 |
by (simp add: * less_le) |
40305
41833242cc42
tuned lemma proposition of properties_for_sort_key
haftmann
parents:
40303
diff
changeset
|
3018 |
next |
46730 | 3019 |
case greater |
3020 |
then have "f l \<noteq> ?pivot" and "\<not> f l < ?pivot" by auto |
|
3021 |
with greater show ?thesis |
|
40346 | 3022 |
by (simp add: filter_sort [symmetric] ** ***) |
40306 | 3023 |
qed |
40303
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
3024 |
qed |
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
3025 |
|
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
haftmann
parents:
40250
diff
changeset
|
3026 |
lemma sort_by_quicksort: |
68386 | 3027 |
"sort xs = sort [x\<leftarrow>xs. x < xs ! (length xs div 2)] |
3028 |
@ [x\<leftarrow>xs. x = xs ! (length xs div 2)] |
|
3029 |
@ 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
|
3030 |
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
|
3031 |
|
68983 | 3032 |
text \<open>A stable parameterized quicksort\<close> |
40347
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3033 |
|
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3034 |
definition part :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'b list \<times> 'b list \<times> 'b list" where |
68386 | 3035 |
"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
|
3036 |
|
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3037 |
lemma part_code [code]: |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3038 |
"part f pivot [] = ([], [], [])" |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3039 |
"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
|
3040 |
if x' < pivot then (x # lts, eqs, gts) |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3041 |
else if x' > pivot then (lts, eqs, x # gts) |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3042 |
else (lts, x # eqs, gts))" |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3043 |
by (auto simp add: part_def Let_def split_def) |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3044 |
|
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3045 |
lemma sort_key_by_quicksort_code [code]: |
60606 | 3046 |
"sort_key f xs = |
3047 |
(case xs of |
|
3048 |
[] \<Rightarrow> [] |
|
40347
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3049 |
| [x] \<Rightarrow> xs |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3050 |
| [x, y] \<Rightarrow> (if f x \<le> f y then xs else [y, x]) |
60606 | 3051 |
| _ \<Rightarrow> |
3052 |
let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs |
|
3053 |
in sort_key f lts @ eqs @ sort_key f gts)" |
|
40347
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3054 |
proof (cases xs) |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3055 |
case Nil then show ?thesis by simp |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3056 |
next |
46921 | 3057 |
case (Cons _ ys) note hyps = Cons show ?thesis |
3058 |
proof (cases ys) |
|
40347
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3059 |
case Nil with hyps show ?thesis by simp |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3060 |
next |
46921 | 3061 |
case (Cons _ zs) note hyps = hyps Cons show ?thesis |
3062 |
proof (cases zs) |
|
40347
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3063 |
case Nil with hyps show ?thesis by auto |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3064 |
next |
58425 | 3065 |
case Cons |
40347
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3066 |
from sort_key_by_quicksort [of f xs] |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3067 |
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
|
3068 |
in sort_key f lts @ eqs @ sort_key f gts)" |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3069 |
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
|
3070 |
with hyps Cons show ?thesis by (simp only: list.cases) |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3071 |
qed |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3072 |
qed |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3073 |
qed |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3074 |
|
39533
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
3075 |
end |
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
haftmann
parents:
39314
diff
changeset
|
3076 |
|
40347
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3077 |
hide_const (open) part |
429bf4388b2f
added code lemmas for stable parametrized quicksort
haftmann
parents:
40346
diff
changeset
|
3078 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
3079 |
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
|
3080 |
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
|
3081 |
|
60515 | 3082 |
lemma mset_update: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3083 |
"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
|
3084 |
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
|
3085 |
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
|
3086 |
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
|
3087 |
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
|
3088 |
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
|
3089 |
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
|
3090 |
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
|
3091 |
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
|
3092 |
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
|
3093 |
with Cons show ?thesis |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3094 |
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
|
3095 |
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
|
3096 |
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
|
3097 |
|
60515 | 3098 |
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
|
3099 |
"i < length ls \<Longrightarrow> j < length ls \<Longrightarrow> |
60515 | 3100 |
mset (ls[j := ls ! i, i := ls ! j]) = mset ls" |
3101 |
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
|
3102 |
|
73327
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
haftmann
parents:
73301
diff
changeset
|
3103 |
lemma mset_eq_finite: |
73301 | 3104 |
\<open>finite {ys. mset ys = mset xs}\<close> |
3105 |
proof - |
|
3106 |
have \<open>{ys. mset ys = mset xs} \<subseteq> {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close> |
|
3107 |
by (auto simp add: dest: mset_eq_setD mset_eq_length) |
|
3108 |
moreover have \<open>finite {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close> |
|
3109 |
using finite_lists_length_le by blast |
|
3110 |
ultimately show ?thesis |
|
3111 |
by (rule finite_subset) |
|
3112 |
qed |
|
3113 |
||
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
|
3114 |
|
60500 | 3115 |
subsection \<open>The multiset order\<close> |
3116 |
||
60606 | 3117 |
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
|
3118 |
"mult1 r = {(N, M). \<exists>a M0 K. M = add_mset a M0 \<and> N = M0 + K \<and> |
60607 | 3119 |
(\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r)}" |
60606 | 3120 |
|
3121 |
definition mult :: "('a \<times> 'a) set \<Rightarrow> ('a multiset \<times> 'a multiset) set" where |
|
37765 | 3122 |
"mult r = (mult1 r)\<^sup>+" |
10249 | 3123 |
|
74858
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
desharna
parents:
74806
diff
changeset
|
3124 |
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
|
3125 |
"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
|
3126 |
|
76589
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3127 |
declare multp_def[pred_set_conv] |
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3128 |
|
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
3129 |
lemma mult1I: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3130 |
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
|
3131 |
shows "(N, M) \<in> mult1 r" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
3132 |
using assms unfolding mult1_def by blast |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
3133 |
|
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
3134 |
lemma mult1E: |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
3135 |
assumes "(N, M) \<in> mult1 r" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3136 |
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
|
3137 |
using assms unfolding mult1_def by blast |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
3138 |
|
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3139 |
lemma mono_mult1: |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3140 |
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
|
3141 |
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
|
3142 |
|
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3143 |
lemma mono_mult: |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3144 |
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
|
3145 |
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
|
3146 |
|
74859 | 3147 |
lemma mono_multp[mono]: "r \<le> r' \<Longrightarrow> multp r \<le> multp r'" |
3148 |
unfolding le_fun_def le_bool_def |
|
3149 |
proof (intro allI impI) |
|
3150 |
fix M N :: "'a multiset" |
|
3151 |
assume "\<forall>x xa. r x xa \<longrightarrow> r' x xa" |
|
3152 |
hence "{(x, y). r x y} \<subseteq> {(x, y). r' x y}" |
|
3153 |
by blast |
|
3154 |
thus "multp r M N \<Longrightarrow> multp r' M N" |
|
3155 |
unfolding multp_def |
|
3156 |
by (fact mono_mult[THEN subsetD, rotated]) |
|
3157 |
qed |
|
3158 |
||
23751 | 3159 |
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
|
3160 |
by (simp add: mult1_def) |
10249 | 3161 |
|
74860 | 3162 |
|
3163 |
subsubsection \<open>Well-foundedness\<close> |
|
3164 |
||
60608 | 3165 |
lemma less_add: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3166 |
assumes mult1: "(N, add_mset a M0) \<in> mult1 r" |
60608 | 3167 |
shows |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3168 |
"(\<exists>M. (M, M0) \<in> mult1 r \<and> N = add_mset a M) \<or> |
60608 | 3169 |
(\<exists>K. (\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r) \<and> N = M0 + K)" |
3170 |
proof - |
|
60607 | 3171 |
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
|
3172 |
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
|
3173 |
obtain a' M0' K where M0: "add_mset a M0 = add_mset a' M0'" |
60608 | 3174 |
and N: "N = M0' + K" |
3175 |
and r: "?r K a'" |
|
3176 |
using mult1 unfolding mult1_def by auto |
|
3177 |
show ?thesis (is "?case1 \<or> ?case2") |
|
60606 | 3178 |
proof - |
3179 |
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
|
3180 |
| K' where "M0 = add_mset a' K'" "M0' = add_mset a K'" |
60606 | 3181 |
by atomize_elim (simp only: add_eq_conv_ex) |
18258 | 3182 |
then show ?thesis |
60606 | 3183 |
proof cases |
3184 |
case 1 |
|
11464 | 3185 |
with N r have "?r K a \<and> N = M0 + K" by simp |
60606 | 3186 |
then have ?case2 .. |
3187 |
then show ?thesis .. |
|
10249 | 3188 |
next |
60606 | 3189 |
case 2 |
63794
bcec0534aeea
clean argument of simp add
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63793
diff
changeset
|
3190 |
from N 2(2) have n: "N = add_mset a (K' + K)" by simp |
60606 | 3191 |
with r 2(1) have "?R (K' + K) M0" by blast |
60608 | 3192 |
with n have ?case1 by (simp add: mult1_def) |
60606 | 3193 |
then show ?thesis .. |
10249 | 3194 |
qed |
3195 |
qed |
|
3196 |
qed |
|
3197 |
||
60608 | 3198 |
lemma all_accessible: |
3199 |
assumes "wf r" |
|
3200 |
shows "\<forall>M. M \<in> Wellfounded.acc (mult1 r)" |
|
10249 | 3201 |
proof |
3202 |
let ?R = "mult1 r" |
|
54295 | 3203 |
let ?W = "Wellfounded.acc ?R" |
10249 | 3204 |
{ |
3205 |
fix M M0 a |
|
23751 | 3206 |
assume M0: "M0 \<in> ?W" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3207 |
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
|
3208 |
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
|
3209 |
have "add_mset a M0 \<in> ?W" |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3210 |
proof (rule accI [of "add_mset a M0"]) |
10249 | 3211 |
fix N |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3212 |
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
|
3213 |
then consider M where "(M, M0) \<in> ?R" "N = add_mset a M" |
60608 | 3214 |
| K where "\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r" "N = M0 + K" |
3215 |
by atomize_elim (rule less_add) |
|
23751 | 3216 |
then show "N \<in> ?W" |
60608 | 3217 |
proof cases |
3218 |
case 1 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3219 |
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
|
3220 |
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
|
3221 |
then show "N \<in> ?W" by (simp only: \<open>N = add_mset a M\<close>) |
10249 | 3222 |
next |
60608 | 3223 |
case 2 |
3224 |
from this(1) have "M0 + K \<in> ?W" |
|
10249 | 3225 |
proof (induct K) |
18730 | 3226 |
case empty |
23751 | 3227 |
from M0 show "M0 + {#} \<in> ?W" by simp |
18730 | 3228 |
next |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3229 |
case (add x K) |
23751 | 3230 |
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
|
3231 |
with wf_hyp have "\<forall>M \<in> ?W. add_mset x M \<in> ?W" by blast |
23751 | 3232 |
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
|
3233 |
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
|
3234 |
then show "M0 + (add_mset x K) \<in> ?W" by simp |
10249 | 3235 |
qed |
60608 | 3236 |
then show "N \<in> ?W" by (simp only: 2(2)) |
10249 | 3237 |
qed |
3238 |
qed |
|
3239 |
} note tedious_reasoning = this |
|
3240 |
||
60608 | 3241 |
show "M \<in> ?W" for M |
10249 | 3242 |
proof (induct M) |
23751 | 3243 |
show "{#} \<in> ?W" |
10249 | 3244 |
proof (rule accI) |
23751 | 3245 |
fix b assume "(b, {#}) \<in> ?R" |
3246 |
with not_less_empty show "b \<in> ?W" by contradiction |
|
10249 | 3247 |
qed |
3248 |
||
23751 | 3249 |
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
|
3250 |
from \<open>wf r\<close> have "\<forall>M \<in> ?W. add_mset a M \<in> ?W" |
10249 | 3251 |
proof induct |
3252 |
fix a |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3253 |
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
|
3254 |
show "\<forall>M \<in> ?W. add_mset a M \<in> ?W" |
10249 | 3255 |
proof |
23751 | 3256 |
fix M assume "M \<in> ?W" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3257 |
then show "add_mset a M \<in> ?W" |
23373 | 3258 |
by (rule acc_induct) (rule tedious_reasoning [OF _ r]) |
10249 | 3259 |
qed |
3260 |
qed |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3261 |
from this and \<open>M \<in> ?W\<close> show "add_mset a M \<in> ?W" .. |
10249 | 3262 |
qed |
3263 |
qed |
|
3264 |
||
74860 | 3265 |
lemma wf_mult1: "wf r \<Longrightarrow> wf (mult1 r)" |
3266 |
by (rule acc_wfI) (rule all_accessible) |
|
3267 |
||
3268 |
lemma wf_mult: "wf r \<Longrightarrow> wf (mult r)" |
|
3269 |
unfolding mult_def by (rule wf_trancl) (rule wf_mult1) |
|
3270 |
||
3271 |
lemma wfP_multp: "wfP r \<Longrightarrow> wfP (multp r)" |
|
3272 |
unfolding multp_def wfP_def |
|
3273 |
by (simp add: wf_mult) |
|
10249 | 3274 |
|
3275 |
||
60500 | 3276 |
subsubsection \<open>Closure-free presentation\<close> |
3277 |
||
3278 |
text \<open>One direction.\<close> |
|
10249 | 3279 |
lemma mult_implies_one_step: |
63795
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3280 |
assumes |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3281 |
trans: "trans r" and |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3282 |
MN: "(M, N) \<in> mult r" |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3283 |
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
|
3284 |
using MN unfolding mult_def mult1_def |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3285 |
proof (induction rule: converse_trancl_induct) |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3286 |
case (base y) |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3287 |
then show ?case by force |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3288 |
next |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3289 |
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
|
3290 |
obtain I J K where |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3291 |
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
|
3292 |
using N_decomp by blast |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3293 |
obtain a M0 K' where |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3294 |
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
|
3295 |
using yz by blast |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3296 |
show ?case |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3297 |
proof (cases "a \<in># K") |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3298 |
case True |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3299 |
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
|
3300 |
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
|
3301 |
ultimately show ?thesis |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3302 |
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
|
3303 |
(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
|
3304 |
next |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3305 |
case False |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3306 |
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
|
3307 |
moreover have "M0 = I + K - {#a#}" |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3308 |
using N(2) z by force |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3309 |
ultimately show ?thesis |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3310 |
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
|
3311 |
rule_tac x = "K + K'" in exI) |
64017
6e7bf7678518
more multiset simp rules
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63924
diff
changeset
|
3312 |
(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
|
3313 |
qed |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3314 |
qed |
10249 | 3315 |
|
76589
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3316 |
lemma multp_implies_one_step: |
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3317 |
"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
|
3318 |
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
|
3319 |
|
17161 | 3320 |
lemma one_step_implies_mult: |
63795
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3321 |
assumes |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3322 |
"J \<noteq> {#}" and |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3323 |
"\<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
|
3324 |
shows "(I + K, I + J) \<in> mult r" |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3325 |
using assms |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3326 |
proof (induction "size J" arbitrary: I J K) |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3327 |
case 0 |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3328 |
then show ?case by auto |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3329 |
next |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3330 |
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
|
3331 |
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
|
3332 |
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
|
3333 |
show ?case |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3334 |
proof (cases "J' = {#}") |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3335 |
case True |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3336 |
then show ?thesis |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3337 |
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
|
3338 |
next |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3339 |
case [simp]: False |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3340 |
have K: "K = {#x \<in># K. (x, a) \<in> r#} + {#x \<in># K. (x, a) \<notin> r#}" |
68992 | 3341 |
by simp |
63795
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3342 |
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
|
3343 |
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
|
3344 |
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
|
3345 |
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
|
3346 |
by (fastforce simp: J mult1_def mult_def) |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3347 |
ultimately show ?thesis |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3348 |
unfolding mult_def by simp |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3349 |
qed |
7f6128adfe67
tuning multisets; more interpretations
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63794
diff
changeset
|
3350 |
qed |
10249 | 3351 |
|
76589
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3352 |
lemma one_step_implies_multp: |
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3353 |
"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
|
3354 |
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
|
3355 |
|
65047 | 3356 |
lemma subset_implies_mult: |
3357 |
assumes sub: "A \<subset># B" |
|
3358 |
shows "(A, B) \<in> mult r" |
|
3359 |
proof - |
|
3360 |
have ApBmA: "A + (B - A) = B" |
|
3361 |
using sub by simp |
|
3362 |
have BmA: "B - A \<noteq> {#}" |
|
3363 |
using sub by (simp add: Diff_eq_empty_iff_mset subset_mset.less_le_not_le) |
|
3364 |
thus ?thesis |
|
3365 |
by (rule one_step_implies_mult[of "B - A" "{#}" _ A, unfolded ApBmA, simplified]) |
|
3366 |
qed |
|
3367 |
||
76589
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3368 |
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
|
3369 |
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
|
3370 |
|
77688 | 3371 |
lemma multp_repeat_mset_repeat_msetI: |
3372 |
assumes "transp R" and "multp R A B" and "n \<noteq> 0" |
|
3373 |
shows "multp R (repeat_mset n A) (repeat_mset n B)" |
|
3374 |
proof - |
|
3375 |
from \<open>transp R\<close> \<open>multp R A B\<close> obtain I J K where |
|
3376 |
"B = I + J" and "A = I + K" and "J \<noteq> {#}" and "\<forall>k \<in># K. \<exists>x \<in># J. R k x" |
|
3377 |
by (auto dest: multp_implies_one_step) |
|
3378 |
||
3379 |
have repeat_n_A_eq: "repeat_mset n A = repeat_mset n I + repeat_mset n K" |
|
3380 |
using \<open>A = I + K\<close> by simp |
|
3381 |
||
3382 |
have repeat_n_B_eq: "repeat_mset n B = repeat_mset n I + repeat_mset n J" |
|
3383 |
using \<open>B = I + J\<close> by simp |
|
3384 |
||
3385 |
show ?thesis |
|
3386 |
unfolding repeat_n_A_eq repeat_n_B_eq |
|
3387 |
proof (rule one_step_implies_multp) |
|
3388 |
from \<open>n \<noteq> 0\<close> show "repeat_mset n J \<noteq> {#}" |
|
3389 |
using \<open>J \<noteq> {#}\<close> |
|
3390 |
by (simp add: repeat_mset_eq_empty_iff) |
|
3391 |
next |
|
3392 |
show "\<forall>k \<in># repeat_mset n K. \<exists>j \<in># repeat_mset n J. R k j" |
|
3393 |
using \<open>\<forall>k \<in># K. \<exists>x \<in># J. R k x\<close> |
|
3394 |
by (metis count_greater_zero_iff nat_0_less_mult_iff repeat_mset.rep_eq) |
|
3395 |
qed |
|
3396 |
qed |
|
3397 |
||
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3398 |
|
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
|
3399 |
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
|
3400 |
|
76401
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3401 |
lemma multp_mono_strong: |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3402 |
assumes "multp R M1 M2" and "transp R" and |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3403 |
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
|
3404 |
shows "multp S M1 M2" |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3405 |
proof - |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3406 |
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
|
3407 |
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
|
3408 |
show ?thesis |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3409 |
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
|
3410 |
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
|
3411 |
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
|
3412 |
using S_if_R |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3413 |
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
|
3414 |
qed |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3415 |
qed |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3416 |
|
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3417 |
lemma mult_mono_strong: |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3418 |
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
|
3419 |
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
|
3420 |
shows "(M1, M2) \<in> mult s" |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3421 |
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
|
3422 |
unfolded multp_def transp_trans_eq, simplified] |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3423 |
by blast |
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
desharna
parents:
76359
diff
changeset
|
3424 |
|
75584
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3425 |
lemma monotone_on_multp_multp_image_mset: |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3426 |
assumes "monotone_on A orda ordb f" and "transp orda" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3427 |
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
|
3428 |
proof (rule monotone_onI) |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3429 |
fix M1 M2 |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3430 |
assume |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3431 |
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
|
3432 |
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
|
3433 |
M1_lt_M2: "multp orda M1 M2" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3434 |
|
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3435 |
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
|
3436 |
M2_eq: "M2 = I + J" and |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3437 |
M1_eq: "M1 = I + K" and |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3438 |
J_neq_mempty: "J \<noteq> {#}" and |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3439 |
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
|
3440 |
by metis |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3441 |
|
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3442 |
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
|
3443 |
proof (intro one_step_implies_multp ballI) |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3444 |
show "image_mset f J \<noteq> {#}" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3445 |
using J_neq_mempty by simp |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3446 |
next |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3447 |
fix k' assume "k'\<in>#image_mset f K" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3448 |
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
|
3449 |
by auto |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3450 |
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
|
3451 |
using ball_K_less by auto |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3452 |
|
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3453 |
have "ordb (f k) (f j)" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3454 |
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
|
3455 |
show "k \<in> A" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3456 |
using M1_eq M1_in k_in by auto |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3457 |
next |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3458 |
show "j \<in> A" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3459 |
using M2_eq M2_in j_in by auto |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3460 |
qed |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3461 |
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
|
3462 |
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
|
3463 |
qed |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3464 |
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
|
3465 |
by (simp add: M1_eq M2_eq) |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3466 |
qed |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3467 |
|
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3468 |
lemma monotone_multp_multp_image_mset: |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3469 |
assumes "monotone orda ordb f" and "transp orda" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3470 |
shows "monotone (multp orda) (multp ordb) (image_mset f)" |
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
desharna
parents:
75560
diff
changeset
|
3471 |
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
|
3472 |
|
77832 | 3473 |
lemma multp_image_mset_image_msetI: |
3474 |
assumes "multp (\<lambda>x y. R (f x) (f y)) M1 M2" and "transp R" |
|
3475 |
shows "multp R (image_mset f M1) (image_mset f M2)" |
|
3476 |
proof - |
|
3477 |
from \<open>transp R\<close> have "transp (\<lambda>x y. R (f x) (f y))" |
|
3478 |
by (auto intro: transpI dest: transpD) |
|
3479 |
with \<open>multp (\<lambda>x y. R (f x) (f y)) M1 M2\<close> obtain I J K where |
|
3480 |
"M2 = I + J" and "M1 = I + K" and "J \<noteq> {#}" and "\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)" |
|
3481 |
using multp_implies_one_step by blast |
|
3482 |
||
3483 |
have "multp R (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)" |
|
3484 |
proof (rule one_step_implies_multp) |
|
3485 |
show "image_mset f J \<noteq> {#}" |
|
3486 |
by (simp add: \<open>J \<noteq> {#}\<close>) |
|
3487 |
next |
|
3488 |
show "\<forall>k\<in>#image_mset f K. \<exists>j\<in>#image_mset f J. R k j" |
|
3489 |
by (simp add: \<open>\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)\<close>) |
|
3490 |
qed |
|
3491 |
thus ?thesis |
|
3492 |
by (simp add: \<open>M1 = I + K\<close> \<open>M2 = I + J\<close>) |
|
3493 |
qed |
|
3494 |
||
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
|
3495 |
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
|
3496 |
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
|
3497 |
"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
|
3498 |
"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
|
3499 |
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
|
3500 |
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
|
3501 |
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
|
3502 |
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
|
3503 |
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
|
3504 |
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
|
3505 |
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
|
3506 |
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
|
3507 |
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
|
3508 |
|
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
|
3509 |
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
|
3510 |
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
|
3511 |
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
|
3512 |
|
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
|
3513 |
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
|
3514 |
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
|
3515 |
|
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
|
3516 |
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
|
3517 |
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
|
3518 |
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
|
3519 |
|
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
|
3520 |
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
|
3521 |
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
|
3522 |
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
|
3523 |
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
|
3524 |
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
|
3525 |
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
|
3526 |
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
|
3527 |
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
|
3528 |
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
|
3529 |
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
|
3530 |
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
|
3531 |
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
|
3532 |
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
|
3533 |
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
|
3534 |
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
|
3535 |
|
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
|
3536 |
|
74862
aa51e974b688
added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
desharna
parents:
74861
diff
changeset
|
3537 |
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
|
3538 |
|
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3539 |
lemma mult_cancel: |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3540 |
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
|
3541 |
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
|
3542 |
proof |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3543 |
assume ?L thus ?R |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3544 |
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
|
3545 |
proof (induct Z) |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3546 |
case (add z Z) |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
3547 |
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
|
3548 |
"\<forall>x \<in> set_mset X'. \<exists>y \<in> set_mset Y'. (x, y) \<in> s" |
64911 | 3549 |
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
|
3550 |
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
|
3551 |
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
|
3552 |
thus ?case |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3553 |
proof (cases) |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3554 |
case 1 thus ?thesis |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3555 |
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
|
3556 |
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
|
3557 |
next |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3558 |
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
|
3559 |
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
|
3560 |
by (auto simp: irrefl_on_def) |
64911 | 3561 |
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
|
3562 |
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
|
3563 |
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
|
3564 |
ultimately show ?thesis |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3565 |
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
|
3566 |
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
|
3567 |
elim: irrefl_on_subset) |
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3568 |
qed |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3569 |
qed auto |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3570 |
next |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3571 |
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
|
3572 |
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 | 3573 |
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
|
3574 |
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
|
3575 |
qed |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3576 |
|
76589
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3577 |
lemma multp_cancel: |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3578 |
"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
|
3579 |
by (rule mult_cancel[to_pred]) |
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3580 |
|
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3581 |
lemma mult_cancel_add_mset: |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3582 |
"trans r \<Longrightarrow> irrefl_on {z} r \<Longrightarrow> |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3583 |
((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
|
3584 |
by (rule mult_cancel[of _ "{#_#}", simplified]) |
76589
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3585 |
|
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3586 |
lemma multp_cancel_add_mset: |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3587 |
"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
|
3588 |
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
|
3589 |
|
74804
5749fefd3fa0
simplified mult_cancel_max and introduced orginal lemma as mult_cancel_max0
desharna
parents:
74803
diff
changeset
|
3590 |
lemma mult_cancel_max0: |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3591 |
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
|
3592 |
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
|
3593 |
proof - |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3594 |
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
|
3595 |
proof (rule mult_cancel) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3596 |
from assms show "trans s" |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3597 |
by simp |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3598 |
next |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3599 |
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
|
3600 |
by simp |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3601 |
qed |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3602 |
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
|
3603 |
by (auto simp flip: count_inject) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3604 |
ultimately show ?thesis |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3605 |
by simp |
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3606 |
qed |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3607 |
|
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3608 |
lemma mult_cancel_max: |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3609 |
"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
|
3610 |
(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
|
3611 |
by (rule mult_cancel_max0[simplified]) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3612 |
|
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3613 |
lemma multp_cancel_max: |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3614 |
"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
|
3615 |
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
|
3616 |
|
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3617 |
|
77049
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3618 |
subsubsection \<open>Strict partial-order properties\<close> |
74864 | 3619 |
|
3620 |
lemma mult1_lessE: |
|
3621 |
assumes "(N, M) \<in> mult1 {(a, b). r a b}" and "asymp r" |
|
3622 |
obtains a M0 K where "M = add_mset a M0" "N = M0 + K" |
|
3623 |
"a \<notin># K" "\<And>b. b \<in># K \<Longrightarrow> r b a" |
|
3624 |
proof - |
|
3625 |
from assms obtain a M0 K where "M = add_mset a M0" "N = M0 + K" and |
|
3626 |
*: "b \<in># K \<Longrightarrow> r b a" for b by (blast elim: mult1E) |
|
3627 |
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
|
3628 |
using \<open>asymp r\<close> by (meson asympD) |
74864 | 3629 |
ultimately show thesis by (auto intro: that) |
3630 |
qed |
|
3631 |
||
79575
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3632 |
lemma trans_on_mult: |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3633 |
assumes "trans_on A r" and "\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A" |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3634 |
shows "trans_on B (mult r)" |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3635 |
using assms by (metis mult_def subset_UNIV trans_on_subset trans_trancl) |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3636 |
|
74865
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3637 |
lemma trans_mult: "trans r \<Longrightarrow> trans (mult r)" |
79575
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3638 |
using trans_on_mult[of UNIV r UNIV, simplified] . |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3639 |
|
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3640 |
lemma transp_on_multp: |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3641 |
assumes "transp_on A r" and "\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A" |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3642 |
shows "transp_on B (multp r)" |
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3643 |
by (metis mult_def multp_def transD trans_trancl transp_onI) |
74865
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3644 |
|
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3645 |
lemma transp_multp: "transp r \<Longrightarrow> transp (multp r)" |
79575
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
desharna
parents:
78099
diff
changeset
|
3646 |
using transp_on_multp[of UNIV r UNIV, simplified] . |
74865
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3647 |
|
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3648 |
lemma irrefl_mult: |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3649 |
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
|
3650 |
shows "irrefl (mult r)" |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3651 |
proof (intro irreflI notI) |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3652 |
fix M |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3653 |
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
|
3654 |
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
|
3655 |
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
|
3656 |
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
|
3657 |
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
|
3658 |
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
|
3659 |
hence "set_mset K = {}" |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3660 |
using ** |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3661 |
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
|
3662 |
case empty |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3663 |
thus ?case by simp |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3664 |
next |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3665 |
case (insert x F) |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3666 |
have False |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3667 |
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
|
3668 |
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
|
3669 |
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
|
3670 |
thus ?case .. |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3671 |
qed |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3672 |
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
|
3673 |
qed |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3674 |
|
76589
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
desharna
parents:
76570
diff
changeset
|
3675 |
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
|
3676 |
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
|
3677 |
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
|
3678 |
|
74864 | 3679 |
instantiation multiset :: (preorder) order begin |
3680 |
||
3681 |
definition less_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" |
|
3682 |
where "M < N \<longleftrightarrow> multp (<) M N" |
|
3683 |
||
3684 |
definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" |
|
3685 |
where "less_eq_multiset M N \<longleftrightarrow> M < N \<or> M = N" |
|
3686 |
||
3687 |
instance |
|
74865
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3688 |
proof intro_classes |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3689 |
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
|
3690 |
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
|
3691 |
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
|
3692 |
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
|
3693 |
next |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3694 |
fix M :: "'a multiset" |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3695 |
show "M \<le> M" |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3696 |
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
|
3697 |
by simp |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3698 |
next |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3699 |
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
|
3700 |
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
|
3701 |
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
|
3702 |
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
|
3703 |
by blast |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3704 |
next |
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
desharna
parents:
74864
diff
changeset
|
3705 |
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
|
3706 |
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
|
3707 |
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
|
3708 |
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
|
3709 |
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
|
3710 |
by blast |
74864 | 3711 |
qed |
3712 |
||
3713 |
end |
|
3714 |
||
3715 |
lemma mset_le_irrefl [elim!]: |
|
3716 |
fixes M :: "'a::preorder multiset" |
|
3717 |
shows "M < M \<Longrightarrow> R" |
|
3718 |
by simp |
|
3719 |
||
74868
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
desharna
parents:
74865
diff
changeset
|
3720 |
lemma wfP_less_multiset[simp]: |
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
desharna
parents:
74865
diff
changeset
|
3721 |
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
|
3722 |
shows "wfP ((<) :: 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool)" |
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
desharna
parents:
74865
diff
changeset
|
3723 |
using wfP_multp[OF wfP_less] less_multiset_def |
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
desharna
parents:
74865
diff
changeset
|
3724 |
by (metis wfPUNIVI wfP_induct) |
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
desharna
parents:
74865
diff
changeset
|
3725 |
|
74864 | 3726 |
|
77049
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3727 |
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
|
3728 |
|
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3729 |
lemma total_on_mult: |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3730 |
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
|
3731 |
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
|
3732 |
proof (rule total_onI) |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3733 |
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
|
3734 |
let ?I = "M1 \<inter># M2" |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3735 |
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
|
3736 |
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
|
3737 |
case True |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3738 |
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
|
3739 |
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
|
3740 |
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
|
3741 |
next |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3742 |
case False |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3743 |
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
|
3744 |
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
|
3745 |
with False obtain greatest1 where |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3746 |
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
|
3747 |
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
|
3748 |
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
|
3749 |
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
|
3750 |
|
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3751 |
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
|
3752 |
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
|
3753 |
with False obtain greatest2 where |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3754 |
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
|
3755 |
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
|
3756 |
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
|
3757 |
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
|
3758 |
|
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3759 |
have "greatest1 \<noteq> greatest2" |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3760 |
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
|
3761 |
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
|
3762 |
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
|
3763 |
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
|
3764 |
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
|
3765 |
\<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
|
3766 |
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
|
3767 |
thus ?thesis |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3768 |
proof (elim disjE) |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3769 |
assume "(greatest1, greatest2) \<in> r" |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3770 |
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
|
3771 |
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
|
3772 |
show "M2 - ?I \<noteq> {#}" |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3773 |
using False by force |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3774 |
next |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3775 |
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
|
3776 |
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
|
3777 |
by (metis assms(2) transD) |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3778 |
qed |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3779 |
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
|
3780 |
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
|
3781 |
subset_mset.inf.cobounded2) |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3782 |
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
|
3783 |
next |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3784 |
assume "(greatest2, greatest1) \<in> r" |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3785 |
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
|
3786 |
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
|
3787 |
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
|
3788 |
using False by force |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3789 |
next |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3790 |
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
|
3791 |
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
|
3792 |
by (metis assms(2) transD) |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3793 |
qed |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3794 |
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
|
3795 |
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
|
3796 |
subset_mset.inf.cobounded2) |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3797 |
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
|
3798 |
qed |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3799 |
qed |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3800 |
qed |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3801 |
|
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3802 |
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
|
3803 |
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
|
3804 |
|
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3805 |
lemma totalp_on_multp: |
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3806 |
"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
|
3807 |
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
|
3808 |
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
|
3809 |
|
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3810 |
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
|
3811 |
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
|
3812 |
|
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
desharna
parents:
76755
diff
changeset
|
3813 |
|
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3814 |
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
|
3815 |
|
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
haftmann
parents:
63060
diff
changeset
|
3816 |
text \<open> |
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3817 |
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
|
3818 |
executable whenever the given predicate \<open>P\<close> is. Together with the standard |
67398 | 3819 |
code equations for \<open>(\<inter>#\<close>) and \<open>(-\<close>) this should yield quadratic |
74803 | 3820 |
(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
|
3821 |
\<close> |
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3822 |
|
74803 | 3823 |
definition multp_code :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where |
3824 |
"multp_code P N M = |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
3825 |
(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
|
3826 |
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
|
3827 |
|
74803 | 3828 |
definition multeqp_code :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where |
3829 |
"multeqp_code P N M = |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
3830 |
(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
|
3831 |
(\<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
|
3832 |
|
74805 | 3833 |
lemma multp_code_iff_mult: |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3834 |
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
|
3835 |
[simp]: "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R" |
74803 | 3836 |
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
|
3837 |
proof - |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
3838 |
have *: "M \<inter># N + (N - M \<inter># N) = N" "M \<inter># N + (M - M \<inter># N) = M" |
68406 | 3839 |
"(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
|
3840 |
show ?thesis |
63088
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
haftmann
parents:
63060
diff
changeset
|
3841 |
proof |
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3842 |
assume ?L thus ?R |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
3843 |
using one_step_implies_mult[of "M - M \<inter># N" "N - M \<inter># N" R "M \<inter># N"] * |
74803 | 3844 |
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
|
3845 |
next |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
3846 |
{ 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
|
3847 |
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
|
3848 |
} note [dest!] = this |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3849 |
assume ?R thus ?L |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3850 |
using mult_cancel_max |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
3851 |
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
|
3852 |
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
|
3853 |
qed |
63660
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3854 |
qed |
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at>
parents:
63560
diff
changeset
|
3855 |
|
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3856 |
lemma multp_code_iff_multp: |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3857 |
"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
|
3858 |
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
|
3859 |
|
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3860 |
lemma multp_code_eq_multp: |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3861 |
assumes "irreflp R" and "transp R" |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3862 |
shows "multp_code R = multp R" |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3863 |
proof (intro ext) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3864 |
fix M N |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3865 |
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
|
3866 |
proof (rule multp_code_iff_multp) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3867 |
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
|
3868 |
by (auto intro: irreflp_on_subset) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3869 |
next |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3870 |
from assms show "transp R" |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3871 |
by simp |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3872 |
qed |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3873 |
qed |
74863
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
desharna
parents:
74862
diff
changeset
|
3874 |
|
74805 | 3875 |
lemma multeqp_code_iff_reflcl_mult: |
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3876 |
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 | 3877 |
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
|
3878 |
proof - |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
3879 |
{ assume "N \<noteq> M" "M - M \<inter># N = {#}" |
68406 | 3880 |
then obtain y where "count N y \<noteq> count M y" by (auto simp flip: count_inject) |
64911 | 3881 |
then have "\<exists>y. count M y < count N y" using \<open>M - M \<inter># N = {#}\<close> |
68406 | 3882 |
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
|
3883 |
} |
74803 | 3884 |
then have "multeqp_code P N M \<longleftrightarrow> multp_code P N M \<or> N = M" |
3885 |
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
|
3886 |
thus ?thesis |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3887 |
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
|
3888 |
qed |
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
haftmann
parents:
63060
diff
changeset
|
3889 |
|
76611
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3890 |
lemma multeqp_code_iff_reflclp_multp: |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3891 |
"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
|
3892 |
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
|
3893 |
|
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3894 |
lemma multeqp_code_eq_reflclp_multp: |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3895 |
assumes "irreflp R" and "transp R" |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3896 |
shows "multeqp_code R = (multp R)\<^sup>=\<^sup>=" |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3897 |
proof (intro ext) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3898 |
fix M N |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3899 |
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
|
3900 |
proof (rule multeqp_code_iff_reflclp_multp) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3901 |
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
|
3902 |
by (auto intro: irreflp_on_subset) |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3903 |
next |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3904 |
from assms show "transp R" |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3905 |
by simp |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3906 |
qed |
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
desharna
parents:
76589
diff
changeset
|
3907 |
qed |
74863
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
desharna
parents:
74862
diff
changeset
|
3908 |
|
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
desharna
parents:
74862
diff
changeset
|
3909 |
|
60500 | 3910 |
subsubsection \<open>Monotonicity of multiset union\<close> |
10249 | 3911 |
|
60606 | 3912 |
lemma mult1_union: "(B, D) \<in> mult1 r \<Longrightarrow> (C + B, C + D) \<in> mult1 r" |
64076 | 3913 |
by (force simp: mult1_def) |
10249 | 3914 |
|
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
3915 |
lemma union_le_mono2: "B < D \<Longrightarrow> C + B < C + (D::'a::preorder multiset)" |
74864 | 3916 |
apply (unfold less_multiset_def multp_def mult_def) |
26178 | 3917 |
apply (erule trancl_induct) |
40249
cd404ecb9107
Remove unnecessary premise of mult1_union
Lars Noschinski <noschinl@in.tum.de>
parents:
39533
diff
changeset
|
3918 |
apply (blast intro: mult1_union) |
cd404ecb9107
Remove unnecessary premise of mult1_union
Lars Noschinski <noschinl@in.tum.de>
parents:
39533
diff
changeset
|
3919 |
apply (blast intro: mult1_union trancl_trans) |
26178 | 3920 |
done |
10249 | 3921 |
|
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
3922 |
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
|
3923 |
apply (subst add.commute [of B C]) |
cc97b347b301
reduced name variants for assoc and commute on plus and mult
haftmann
parents:
57492
diff
changeset
|
3924 |
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
|
3925 |
apply (erule union_le_mono2) |
26178 | 3926 |
done |
10249 | 3927 |
|
17161 | 3928 |
lemma union_less_mono: |
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
3929 |
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
|
3930 |
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
|
3931 |
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
|
3932 |
|
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
3933 |
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
|
3934 |
begin |
a095acd4cfbf
instantiate multiset with multiset ordering
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63360
diff
changeset
|
3935 |
instance |
a095acd4cfbf
instantiate multiset with multiset ordering
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63360
diff
changeset
|
3936 |
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
|
3937 |
end |
15072 | 3938 |
|
63409
3f3223b90239
moved lemmas and locales around (with minor incompatibilities)
blanchet
parents:
63388
diff
changeset
|
3939 |
|
60500 | 3940 |
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
|
3941 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3942 |
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
|
3943 |
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
|
3944 |
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
|
3945 |
by auto |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3946 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3947 |
definition "ms_strict = mult pair_less" |
37765 | 3948 |
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
|
3949 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3950 |
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
|
3951 |
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
|
3952 |
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
|
3953 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3954 |
lemma smsI: |
60495 | 3955 |
"(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
|
3956 |
unfolding ms_strict_def |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3957 |
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
|
3958 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3959 |
lemma wmsI: |
60495 | 3960 |
"(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
|
3961 |
\<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
|
3962 |
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
|
3963 |
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
|
3964 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3965 |
inductive pw_leq |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3966 |
where |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3967 |
pw_leq_empty: "pw_leq {#} {#}" |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3968 |
| 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
|
3969 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3970 |
lemma pw_leq_lstep: |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3971 |
"(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
|
3972 |
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
|
3973 |
|
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3974 |
lemma pw_leq_split: |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3975 |
assumes "pw_leq X Y" |
60495 | 3976 |
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
|
3977 |
using assms |
60606 | 3978 |
proof induct |
29125
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3979 |
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
|
3980 |
next |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3981 |
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
|
3982 |
then obtain A B Z where |
58425 | 3983 |
[simp]: "X = A + Z" "Y = B + Z" |
60495 | 3984 |
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
|
3985 |
by auto |
60606 | 3986 |
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
|
3987 |
unfolding pair_leq_def by auto |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3988 |
thus ?case |
60606 | 3989 |
proof cases |
3990 |
case [simp]: 1 |
|
3991 |
have "{#x#} + X = A + ({#y#}+Z) \<and> {#y#} + Y = B + ({#y#}+Z) \<and> |
|
3992 |
((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
|
3993 |
by auto |
60606 | 3994 |
thus ?thesis by blast |
29125
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3995 |
next |
60606 | 3996 |
case 2 |
29125
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3997 |
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
|
3998 |
have "{#x#} + X = ?A' + Z" |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
3999 |
"{#y#} + Y = ?B' + Z" |
63794
bcec0534aeea
clean argument of simp add
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63793
diff
changeset
|
4000 |
by auto |
58425 | 4001 |
moreover have |
60495 | 4002 |
"(set_mset ?A', set_mset ?B') \<in> max_strict" |
60606 | 4003 |
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
|
4004 |
by (auto elim!: max_ext.cases) |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4005 |
ultimately show ?thesis by blast |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4006 |
qed |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4007 |
qed |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4008 |
|
58425 | 4009 |
lemma |
29125
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4010 |
assumes pwleq: "pw_leq Z Z'" |
60495 | 4011 |
shows ms_strictI: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_strict" |
60606 | 4012 |
and ms_weakI1: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_weak" |
4013 |
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
|
4014 |
proof - |
58425 | 4015 |
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
|
4016 |
obtain A' B' Z'' |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4017 |
where [simp]: "Z = A' + Z''" "Z' = B' + Z''" |
60495 | 4018 |
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
|
4019 |
by blast |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4020 |
{ |
60495 | 4021 |
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
|
4022 |
from mx_or_empty |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4023 |
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
|
4024 |
proof |
60495 | 4025 |
assume max': "(set_mset A', set_mset B') \<in> max_strict" |
4026 |
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
|
4027 |
by (auto simp: max_strict_def intro: max_ext_additive) |
58425 | 4028 |
thus ?thesis by (rule smsI) |
29125
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4029 |
next |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4030 |
assume [simp]: "A' = {#} \<and> B' = {#}" |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4031 |
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
|
4032 |
qed |
60606 | 4033 |
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
|
4034 |
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
|
4035 |
} |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4036 |
from mx_or_empty |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4037 |
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
|
4038 |
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
|
4039 |
qed |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4040 |
|
39301 | 4041 |
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
|
4042 |
and nonempty_plus: "{# x #} + rs \<noteq> {#}" |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4043 |
and nonempty_single: "{# x #} \<noteq> {#}" |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4044 |
by auto |
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
krauss
parents:
28823
diff
changeset
|
4045 |
|
60500 | 4046 |
setup \<open> |
60606 | 4047 |
let |
74634 | 4048 |
fun msetT T = \<^Type>\<open>multiset T\<close>; |
4049 |
||
4050 |
fun mk_mset T [] = \<^instantiate>\<open>'a = T in term \<open>{#}\<close>\<close> |
|
4051 |
| mk_mset T [x] = \<^instantiate>\<open>'a = T and x in term \<open>{#x#}\<close>\<close> |
|
4052 |
| 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 | 4053 |
|
60752 | 4054 |
fun mset_member_tac ctxt m i = |
60606 | 4055 |
if m <= 0 then |
60752 | 4056 |
resolve_tac ctxt @{thms multi_member_this} i ORELSE |
4057 |
resolve_tac ctxt @{thms multi_member_last} i |
|
60606 | 4058 |
else |
60752 | 4059 |
resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i |
4060 |
||
4061 |
fun mset_nonempty_tac ctxt = |
|
4062 |
resolve_tac ctxt @{thms nonempty_plus} ORELSE' |
|
4063 |
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
|
4064 |
|
60606 | 4065 |
fun regroup_munion_conv ctxt = |
73393 | 4066 |
Function_Lib.regroup_conv ctxt \<^const_abbrev>\<open>empty_mset\<close> \<^const_name>\<open>plus\<close> |
60606 | 4067 |
(map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral})) |
4068 |
||
60752 | 4069 |
fun unfold_pwleq_tac ctxt i = |
4070 |
(resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st)) |
|
4071 |
ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i) |
|
4072 |
ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i) |
|
60606 | 4073 |
|
4074 |
val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union}, |
|
4075 |
@{thm Un_insert_left}, @{thm Un_empty_left}] |
|
4076 |
in |
|
4077 |
ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset |
|
4078 |
{ |
|
4079 |
msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, |
|
4080 |
mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, |
|
4081 |
mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps, |
|
4082 |
smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1}, |
|
60752 | 4083 |
reduction_pair = @{thm ms_reduction_pair} |
60606 | 4084 |
}) |
4085 |
end |
|
60500 | 4086 |
\<close> |
4087 |
||
4088 |
||
4089 |
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
|
4090 |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39301
diff
changeset
|
4091 |
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
|
4092 |
|
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
|
4093 |
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
|
4094 |
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
|
4095 |
|
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
|
4096 |
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
|
4097 |
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
|
4098 |
|
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
|
4099 |
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
|
4100 |
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
|
4101 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4102 |
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
|
4103 |
|
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
|
4104 |
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
|
4105 |
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
|
4106 |
|
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
|
4107 |
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
|
4108 |
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
|
4109 |
|
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
|
4110 |
lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y" |
59557 | 4111 |
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
|
4112 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4113 |
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
|
4114 |
by (fact subset_mset.less_trans) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4115 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
4116 |
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
|
4117 |
by (fact subset_mset.inf.commute) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4118 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
4119 |
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
|
4120 |
by (fact subset_mset.inf.assoc [symmetric]) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4121 |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
4122 |
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
|
4123 |
by (fact subset_mset.inf.left_commute) |
35268
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4124 |
|
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4125 |
lemmas multiset_inter_ac = |
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4126 |
multiset_inter_commute |
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4127 |
multiset_inter_assoc |
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4128 |
multiset_inter_left_commute |
04673275441a
switched notations for pointwise and multiset order
haftmann
parents:
35028
diff
changeset
|
4129 |
|
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
4130 |
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
|
4131 |
by (fact less_irrefl) |
a095acd4cfbf
instantiate multiset with multiset ordering
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63360
diff
changeset
|
4132 |
|
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
4133 |
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
|
4134 |
by (fact less_trans) |
a095acd4cfbf
instantiate multiset with multiset ordering
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63360
diff
changeset
|
4135 |
|
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
4136 |
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
|
4137 |
by (fact less_not_sym) |
a095acd4cfbf
instantiate multiset with multiset ordering
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63360
diff
changeset
|
4138 |
|
63410
9789ccc2a477
more instantiations for multiset
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63409
diff
changeset
|
4139 |
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
|
4140 |
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
|
4141 |
|
60500 | 4142 |
declaration \<open> |
60606 | 4143 |
let |
4144 |
fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') = |
|
4145 |
let |
|
4146 |
val (maybe_opt, ps) = |
|
4147 |
Nitpick_Model.dest_plain_fun t' |
|
67398 | 4148 |
||> (~~) |
60606 | 4149 |
||> map (apsnd (snd o HOLogic.dest_number)) |
4150 |
fun elems_for t = |
|
67398 | 4151 |
(case AList.lookup (=) ps t of |
60606 | 4152 |
SOME n => replicate n t |
4153 |
| NONE => [Const (maybe_name, elem_T --> elem_T) $ t]) |
|
4154 |
in |
|
4155 |
(case maps elems_for (all_values elem_T) @ |
|
61333 | 4156 |
(if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of |
74634 | 4157 |
[] => \<^Const>\<open>Groups.zero T\<close> |
4158 |
| ts => foldl1 (fn (s, t) => \<^Const>\<open>add_mset elem_T for s t\<close>) ts) |
|
60606 | 4159 |
end |
4160 |
| multiset_postproc _ _ _ _ t = t |
|
69593 | 4161 |
in Nitpick_Model.register_term_postprocessor \<^typ>\<open>'a multiset\<close> multiset_postproc end |
60500 | 4162 |
\<close> |
4163 |
||
4164 |
||
4165 |
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
|
4166 |
|
60515 | 4167 |
code_datatype mset |
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4168 |
|
60606 | 4169 |
lemma [code]: "{#} = mset []" |
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4170 |
by simp |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4171 |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4172 |
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
|
4173 |
by simp |
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4174 |
|
63195 | 4175 |
lemma [code]: "Multiset.is_empty (mset xs) \<longleftrightarrow> List.null xs" |
4176 |
by (simp add: Multiset.is_empty_def List.null_def) |
|
4177 |
||
60606 | 4178 |
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
|
4179 |
by simp |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4180 |
|
60606 | 4181 |
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
|
4182 |
by simp |
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4183 |
|
60606 | 4184 |
lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)" |
69442 | 4185 |
by simp |
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4186 |
|
60606 | 4187 |
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
|
4188 |
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
|
4189 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4190 |
lemma [code]: |
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
4191 |
"mset xs \<inter># mset ys = |
60515 | 4192 |
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
|
4193 |
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
|
4194 |
proof - |
60515 | 4195 |
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
|
4196 |
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
|
4197 |
(mset xs \<inter># mset ys) + mset zs" |
51623 | 4198 |
by (induct xs arbitrary: ys) |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
4199 |
(auto simp add: inter_add_right1 inter_add_right2 ac_simps) |
51623 | 4200 |
then show ?thesis by simp |
4201 |
qed |
|
4202 |
||
4203 |
lemma [code]: |
|
63919
9aed2da07200
# after multiset intersection and union symbol
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63908
diff
changeset
|
4204 |
"mset xs \<union># mset ys = |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61378
diff
changeset
|
4205 |
mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" |
51623 | 4206 |
proof - |
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
61378
diff
changeset
|
4207 |
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
|
4208 |
(mset xs \<union># mset ys) + mset zs" |
51623 | 4209 |
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
|
4210 |
then show ?thesis by simp |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4211 |
qed |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4212 |
|
59813 | 4213 |
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
|
4214 |
|
60606 | 4215 |
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
|
4216 |
proof - |
60515 | 4217 |
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
|
4218 |
by (induct xs) simp_all |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4219 |
then show ?thesis by simp |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4220 |
qed |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4221 |
|
60515 | 4222 |
declare set_mset_mset [code] |
4223 |
||
4224 |
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
|
4225 |
|
61585 | 4226 |
lemma [code]: \<comment> \<open>not very efficient, but representation-ignorant!\<close> |
60515 | 4227 |
"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
|
4228 |
apply (cases "finite A") |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4229 |
apply simp_all |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4230 |
apply (induct A rule: finite_induct) |
63794
bcec0534aeea
clean argument of simp add
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63793
diff
changeset
|
4231 |
apply simp_all |
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4232 |
done |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4233 |
|
60515 | 4234 |
declare size_mset [code] |
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4235 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4236 |
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
|
4237 |
"subset_eq_mset_impl [] ys = Some (ys \<noteq> [])" |
67398 | 4238 |
| "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
|
4239 |
None \<Rightarrow> None |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4240 |
| 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
|
4241 |
|
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4242 |
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
|
4243 |
(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
|
4244 |
(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
|
4245 |
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
|
4246 |
case (Nil ys) |
64076 | 4247 |
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
|
4248 |
next |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4249 |
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
|
4250 |
show ?case |
67398 | 4251 |
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
|
4252 |
case None |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4253 |
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
|
4254 |
{ |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
4255 |
assume "mset (x # xs) \<subseteq># mset ys" |
60495 | 4256 |
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
|
4257 |
} note nle = this |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4258 |
moreover |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4259 |
{ |
62430
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
4260 |
assume "mset (x # xs) \<subset># mset ys" |
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
haftmann
parents:
62390
diff
changeset
|
4261 |
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
|
4262 |
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
|
4263 |
} |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4264 |
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
|
4265 |
next |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4266 |
case (Some res) |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4267 |
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
|
4268 |
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
|
4269 |
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
|
4270 |
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
|
4271 |
by auto |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4272 |
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
|
4273 |
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
|
4274 |
unfolding id |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4275 |
using Cons[of "ys1 @ ys2"] |
60397
f8a513fedb31
Renaming multiset operators < ~> <#,...
Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
59999
diff
changeset
|
4276 |
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
|
4277 |
qed |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4278 |
qed |
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4279 |
|
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4280 |
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
|
4281 |
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
|
4282 |
|
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4283 |
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
|
4284 |
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
|
4285 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4286 |
instantiation multiset :: (equal) equal |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4287 |
begin |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4288 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4289 |
definition |
55808
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
nipkow
parents:
55565
diff
changeset
|
4290 |
[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
|
4291 |
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
|
4292 |
unfolding equal_multiset_def |
63310
caaacf37943f
normalising multiset theorem names
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63290
diff
changeset
|
4293 |
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
|
4294 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4295 |
instance |
60678 | 4296 |
by standard (simp add: equal_multiset_def) |
4297 |
||
37169
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
blanchet
parents:
37107
diff
changeset
|
4298 |
end |
49388 | 4299 |
|
66313 | 4300 |
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
|
4301 |
|
63830 | 4302 |
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
|
4303 |
proof - |
63830 | 4304 |
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
|
4305 |
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
|
4306 |
then show ?thesis by simp |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4307 |
qed |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4308 |
|
60500 | 4309 |
text \<open> |
69593 | 4310 |
Exercise for the casual reader: add implementations for \<^term>\<open>(\<le>)\<close> |
4311 |
and \<^term>\<open>(<)\<close> (multiset order). |
|
60500 | 4312 |
\<close> |
4313 |
||
4314 |
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
|
4315 |
|
72607 | 4316 |
context |
4317 |
includes term_syntax |
|
4318 |
begin |
|
4319 |
||
4320 |
definition |
|
61076 | 4321 |
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
|
4322 |
\<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where |
60515 | 4323 |
[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
|
4324 |
|
72607 | 4325 |
end |
4326 |
||
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4327 |
instantiation multiset :: (random) random |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4328 |
begin |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4329 |
|
72581 | 4330 |
context |
4331 |
includes state_combinator_syntax |
|
4332 |
begin |
|
4333 |
||
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4334 |
definition |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4335 |
"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
|
4336 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4337 |
instance .. |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4338 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4339 |
end |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4340 |
|
72581 | 4341 |
end |
51600
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4342 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4343 |
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
|
4344 |
begin |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4345 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4346 |
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
|
4347 |
where |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4348 |
"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
|
4349 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4350 |
instance .. |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4351 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4352 |
end |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4353 |
|
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4354 |
hide_const (open) msetify |
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
haftmann
parents:
51599
diff
changeset
|
4355 |
|
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4356 |
|
60500 | 4357 |
subsection \<open>BNF setup\<close> |
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4358 |
|
57966 | 4359 |
definition rel_mset where |
60515 | 4360 |
"rel_mset R X Y \<longleftrightarrow> (\<exists>xs ys. mset xs = X \<and> mset ys = Y \<and> list_all2 R xs ys)" |
4361 |
||
4362 |
lemma mset_zip_take_Cons_drop_twice: |
|
57966 | 4363 |
assumes "length xs = length ys" "j \<le> length xs" |
60515 | 4364 |
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
|
4365 |
add_mset (x,y) (mset (zip xs ys))" |
60606 | 4366 |
using assms |
57966 | 4367 |
proof (induct xs ys arbitrary: x y j rule: list_induct2) |
4368 |
case Nil |
|
4369 |
thus ?case |
|
4370 |
by simp |
|
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4371 |
next |
57966 | 4372 |
case (Cons x xs y ys) |
4373 |
thus ?case |
|
4374 |
proof (cases "j = 0") |
|
4375 |
case True |
|
4376 |
thus ?thesis |
|
4377 |
by simp |
|
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4378 |
next |
57966 | 4379 |
case False |
4380 |
then obtain k where k: "j = Suc k" |
|
60678 | 4381 |
by (cases j) simp |
57966 | 4382 |
hence "k \<le> length xs" |
4383 |
using Cons.prems by auto |
|
60515 | 4384 |
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
|
4385 |
add_mset (x,y) (mset (zip xs ys))" |
57966 | 4386 |
by (rule Cons.hyps(2)) |
4387 |
thus ?thesis |
|
63794
bcec0534aeea
clean argument of simp add
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63793
diff
changeset
|
4388 |
unfolding k by auto |
58425 | 4389 |
qed |
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4390 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4391 |
|
60515 | 4392 |
lemma ex_mset_zip_left: |
4393 |
assumes "length xs = length ys" "mset xs' = mset xs" |
|
4394 |
shows "\<exists>ys'. length ys' = length xs' \<and> mset (zip xs' ys') = mset (zip xs ys)" |
|
58425 | 4395 |
using assms |
57966 | 4396 |
proof (induct xs ys arbitrary: xs' rule: list_induct2) |
4397 |
case Nil |
|
4398 |
thus ?case |
|
4399 |
by auto |
|
4400 |
next |
|
4401 |
case (Cons x xs y ys xs') |
|
4402 |
obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" |
|
60515 | 4403 |
by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) |
58425 | 4404 |
|
63040 | 4405 |
define xsa where "xsa = take j xs' @ drop (Suc j) xs'" |
60515 | 4406 |
have "mset xs' = {#x#} + mset xsa" |
57966 | 4407 |
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
|
4408 |
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
|
4409 |
append_take_drop_id mset.simps(2) mset_append) |
60515 | 4410 |
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
|
4411 |
by (simp add: Cons.prems) |
57966 | 4412 |
then obtain ysa where |
60515 | 4413 |
len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" |
57966 | 4414 |
using Cons.hyps(2) by blast |
4415 |
||
63040 | 4416 |
define ys' where "ys' = take j ysa @ y # drop j ysa" |
57966 | 4417 |
have xs': "xs' = take j xsa @ x # drop j xsa" |
4418 |
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
|
4419 |
by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons |
60515 | 4420 |
length_drop size_mset) |
57966 | 4421 |
have j_len': "j \<le> length xsa" |
4422 |
using j_len xs' xsa_def |
|
4423 |
by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) |
|
4424 |
have "length ys' = length xs'" |
|
4425 |
unfolding ys'_def using Cons.prems len_a ms_x |
|
60515 | 4426 |
by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) |
4427 |
moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" |
|
57966 | 4428 |
unfolding xs' ys'_def |
60515 | 4429 |
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
|
4430 |
(auto simp: len_a ms_a j_len') |
57966 | 4431 |
ultimately show ?case |
4432 |
by blast |
|
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4433 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4434 |
|
57966 | 4435 |
lemma list_all2_reorder_left_invariance: |
60515 | 4436 |
assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs" |
4437 |
shows "\<exists>ys'. list_all2 R xs' ys' \<and> mset ys' = mset ys" |
|
57966 | 4438 |
proof - |
4439 |
have len: "length xs = length ys" |
|
4440 |
using rel list_all2_conv_all_nth by auto |
|
4441 |
obtain ys' where |
|
60515 | 4442 |
len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" |
4443 |
using len ms_x by (metis ex_mset_zip_left) |
|
57966 | 4444 |
have "list_all2 R xs' ys'" |
60515 | 4445 |
using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) |
4446 |
moreover have "mset ys' = mset ys" |
|
4447 |
using len len' ms_xy map_snd_zip mset_map by metis |
|
57966 | 4448 |
ultimately show ?thesis |
4449 |
by blast |
|
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4450 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4451 |
|
60515 | 4452 |
lemma ex_mset: "\<exists>xs. mset xs = X" |
4453 |
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
|
4454 |
|
73301 | 4455 |
inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool" |
62324 | 4456 |
where |
4457 |
"pred_mset P {#}" |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4458 |
| "\<lbrakk>P a; pred_mset P M\<rbrakk> \<Longrightarrow> pred_mset P (add_mset a M)" |
62324 | 4459 |
|
73301 | 4460 |
lemma pred_mset_iff: \<comment> \<open>TODO: alias for \<^const>\<open>Multiset.Ball\<close>\<close> |
4461 |
\<open>pred_mset P M \<longleftrightarrow> Multiset.Ball M P\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) |
|
4462 |
proof |
|
4463 |
assume ?P |
|
4464 |
then show ?Q by induction simp_all |
|
4465 |
next |
|
4466 |
assume ?Q |
|
4467 |
then show ?P |
|
4468 |
by (induction M) (auto intro: pred_mset.intros) |
|
4469 |
qed |
|
4470 |
||
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4471 |
bnf "'a multiset" |
57966 | 4472 |
map: image_mset |
60495 | 4473 |
sets: set_mset |
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4474 |
bd: natLeq |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4475 |
wits: "{#}" |
57966 | 4476 |
rel: rel_mset |
62324 | 4477 |
pred: pred_mset |
57966 | 4478 |
proof - |
4479 |
show "image_mset id = id" |
|
4480 |
by (rule image_mset.id) |
|
60606 | 4481 |
show "image_mset (g \<circ> f) = image_mset g \<circ> image_mset f" for f g |
59813 | 4482 |
unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality) |
60606 | 4483 |
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
|
4484 |
by (induct X) simp_all |
67398 | 4485 |
show "set_mset \<circ> image_mset f = (`) f \<circ> set_mset" for f |
57966 | 4486 |
by auto |
4487 |
show "card_order natLeq" |
|
4488 |
by (rule natLeq_card_order) |
|
4489 |
show "BNF_Cardinal_Arithmetic.cinfinite natLeq" |
|
4490 |
by (rule natLeq_cinfinite) |
|
75624 | 4491 |
show "regularCard natLeq" |
4492 |
by (rule regularCard_natLeq) |
|
4493 |
show "ordLess2 (card_of (set_mset X)) natLeq" for X |
|
57966 | 4494 |
by transfer |
75624 | 4495 |
(auto simp: finite_iff_ordLess_natLeq[symmetric]) |
60606 | 4496 |
show "rel_mset R OO rel_mset S \<le> rel_mset (R OO S)" for R S |
57966 | 4497 |
unfolding rel_mset_def[abs_def] OO_def |
4498 |
apply clarify |
|
60678 | 4499 |
subgoal for X Z Y xs ys' ys zs |
4500 |
apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys]) |
|
4501 |
apply (auto intro: list_all2_trans) |
|
4502 |
done |
|
60606 | 4503 |
done |
4504 |
show "rel_mset R = |
|
62324 | 4505 |
(\<lambda>x y. \<exists>z. set_mset z \<subseteq> {(x, y). R x y} \<and> |
4506 |
image_mset fst z = x \<and> image_mset snd z = y)" for R |
|
4507 |
unfolding rel_mset_def[abs_def] |
|
57966 | 4508 |
apply (rule ext)+ |
62324 | 4509 |
apply safe |
4510 |
apply (rule_tac x = "mset (zip xs ys)" in exI; |
|
68406 | 4511 |
auto simp: in_set_zip list_all2_iff simp flip: mset_map) |
57966 | 4512 |
apply (rename_tac XY) |
60515 | 4513 |
apply (cut_tac X = XY in ex_mset) |
57966 | 4514 |
apply (erule exE) |
4515 |
apply (rename_tac xys) |
|
4516 |
apply (rule_tac x = "map fst xys" in exI) |
|
60515 | 4517 |
apply (auto simp: mset_map) |
57966 | 4518 |
apply (rule_tac x = "map snd xys" in exI) |
60515 | 4519 |
apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd) |
59997 | 4520 |
done |
60606 | 4521 |
show "z \<in> set_mset {#} \<Longrightarrow> False" for z |
57966 | 4522 |
by auto |
62324 | 4523 |
show "pred_mset P = (\<lambda>x. Ball (set_mset x) P)" for P |
73301 | 4524 |
by (simp add: fun_eq_iff pred_mset_iff) |
57966 | 4525 |
qed |
4526 |
||
73301 | 4527 |
inductive rel_mset' :: \<open>('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset \<Rightarrow> bool\<close> |
60606 | 4528 |
where |
57966 | 4529 |
Zero[intro]: "rel_mset' R {#} {#}" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4530 |
| 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 | 4531 |
|
4532 |
lemma rel_mset_Zero: "rel_mset R {#} {#}" |
|
4533 |
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
|
4534 |
|
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4535 |
declare multiset.count[simp] |
73270 | 4536 |
declare count_Abs_multiset[simp] |
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4537 |
declare multiset.count_inverse[simp] |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4538 |
|
57966 | 4539 |
lemma rel_mset_Plus: |
60606 | 4540 |
assumes ab: "R a b" |
4541 |
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
|
4542 |
shows "rel_mset R (add_mset a M) (add_mset b N)" |
60606 | 4543 |
proof - |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4544 |
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
|
4545 |
add_mset b (image_mset snd y) = image_mset snd ya \<and> |
60606 | 4546 |
set_mset ya \<subseteq> {(x, y). R x y}" |
4547 |
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
|
4548 |
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
|
4549 |
thus ?thesis |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4550 |
using assms |
57966 | 4551 |
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
|
4552 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4553 |
|
60606 | 4554 |
lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \<Longrightarrow> rel_mset R M N" |
60678 | 4555 |
by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus) |
57966 | 4556 |
|
60606 | 4557 |
lemma rel_mset_size: "rel_mset R M N \<Longrightarrow> size M = size N" |
60678 | 4558 |
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
|
4559 |
|
73594 | 4560 |
lemma rel_mset_Zero_iff [simp]: |
4561 |
shows "rel_mset rel {#} Y \<longleftrightarrow> Y = {#}" and "rel_mset rel X {#} \<longleftrightarrow> X = {#}" |
|
4562 |
by (auto simp add: rel_mset_Zero dest: rel_mset_size) |
|
4563 |
||
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4564 |
lemma multiset_induct2[case_names empty addL addR]: |
60678 | 4565 |
assumes empty: "P {#} {#}" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4566 |
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
|
4567 |
and addR: "\<And>a M N. P M N \<Longrightarrow> P M (add_mset a N)" |
60678 | 4568 |
shows "P M N" |
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4569 |
apply(induct N rule: multiset_induct) |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4570 |
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
|
4571 |
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
|
4572 |
done |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4573 |
|
59949 | 4574 |
lemma multiset_induct2_size[consumes 1, case_names empty add]: |
60606 | 4575 |
assumes c: "size M = size N" |
4576 |
and empty: "P {#} {#}" |
|
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4577 |
and add: "\<And>a b M N a b. P M N \<Longrightarrow> P (add_mset a M) (add_mset b N)" |
60606 | 4578 |
shows "P M N" |
60678 | 4579 |
using c |
4580 |
proof (induct M arbitrary: N rule: measure_induct_rule[of size]) |
|
60606 | 4581 |
case (less M) |
4582 |
show ?case |
|
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4583 |
proof(cases "M = {#}") |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4584 |
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
|
4585 |
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
|
4586 |
next |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4587 |
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
|
4588 |
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
|
4589 |
then obtain N1 b where N: "N = add_mset b N1" by (metis multi_nonempty_split) |
59949 | 4590 |
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
|
4591 |
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
|
4592 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4593 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4594 |
|
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4595 |
lemma msed_map_invL: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4596 |
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
|
4597 |
shows "\<exists>N1. N = add_mset (f a) N1 \<and> image_mset f M = N1" |
60606 | 4598 |
proof - |
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4599 |
have "f a \<in># N" |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4600 |
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
|
4601 |
then obtain N1 where N: "N = add_mset (f a) N1" using multi_member_split by metis |
57966 | 4602 |
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
|
4603 |
thus ?thesis using N by blast |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4604 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4605 |
|
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4606 |
lemma msed_map_invR: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4607 |
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
|
4608 |
shows "\<exists>M1 a. M = add_mset a M1 \<and> f a = b \<and> image_mset f M1 = N" |
60606 | 4609 |
proof - |
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4610 |
obtain a where a: "a \<in># M" and fa: "f a = b" |
60606 | 4611 |
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
|
4612 |
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
|
4613 |
then obtain M1 where M: "M = add_mset a M1" using multi_member_split by metis |
57966 | 4614 |
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
|
4615 |
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
|
4616 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4617 |
|
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4618 |
lemma msed_rel_invL: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4619 |
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
|
4620 |
shows "\<exists>N1 b. N = add_mset b N1 \<and> R a b \<and> rel_mset R M N1" |
60606 | 4621 |
proof - |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4622 |
obtain K where KM: "image_mset fst K = add_mset a M" |
60606 | 4623 |
and KN: "image_mset snd K = N" and sK: "set_mset K \<subseteq> {(a, b). R a b}" |
4624 |
using assms |
|
4625 |
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
|
4626 |
obtain K1 ab where K: "K = add_mset ab K1" and a: "fst ab = a" |
60606 | 4627 |
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
|
4628 |
obtain N1 where N: "N = add_mset (snd ab) N1" and K1N1: "image_mset snd K1 = N1" |
60606 | 4629 |
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
|
4630 |
have Rab: "R a (snd ab)" using sK a unfolding K by auto |
57966 | 4631 |
have "rel_mset R M N1" using sK K1M K1N1 |
60606 | 4632 |
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
|
4633 |
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
|
4634 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4635 |
|
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4636 |
lemma msed_rel_invR: |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4637 |
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
|
4638 |
shows "\<exists>M1 a. M = add_mset a M1 \<and> R a b \<and> rel_mset R M1 N" |
60606 | 4639 |
proof - |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4640 |
obtain K where KN: "image_mset snd K = add_mset b N" |
60606 | 4641 |
and KM: "image_mset fst K = M" and sK: "set_mset K \<subseteq> {(a, b). R a b}" |
4642 |
using assms |
|
4643 |
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
|
4644 |
obtain K1 ab where K: "K = add_mset ab K1" and b: "snd ab = b" |
60606 | 4645 |
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
|
4646 |
obtain M1 where M: "M = add_mset (fst ab) M1" and K1M1: "image_mset fst K1 = M1" |
60606 | 4647 |
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
|
4648 |
have Rab: "R (fst ab) b" using sK b unfolding K by auto |
57966 | 4649 |
have "rel_mset R M1 N" using sK K1N K1M1 |
60606 | 4650 |
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
|
4651 |
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
|
4652 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4653 |
|
57966 | 4654 |
lemma rel_mset_imp_rel_mset': |
60606 | 4655 |
assumes "rel_mset R M N" |
4656 |
shows "rel_mset' R M N" |
|
59949 | 4657 |
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
|
4658 |
case (less M) |
59949 | 4659 |
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
|
4660 |
show ?case |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4661 |
proof(cases "M = {#}") |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4662 |
case True hence "N = {#}" using c by simp |
57966 | 4663 |
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
|
4664 |
next |
63793
e68a0b651eb5
add_mset constructor in multisets
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents:
63689
diff
changeset
|
4665 |
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
|
4666 |
obtain N1 b where N: "N = add_mset b N1" and R: "R a b" and ms: "rel_mset R M1 N1" |
60606 | 4667 |
using msed_rel_invL[OF less.prems[unfolded M]] by auto |
57966 | 4668 |
have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp |
4669 |
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
|
4670 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4671 |
qed |
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4672 |
|
60606 | 4673 |
lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N" |
60678 | 4674 |
using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto |
57966 | 4675 |
|
69593 | 4676 |
text \<open>The main end product for \<^const>\<open>rel_mset\<close>: inductive characterization:\<close> |
61337 | 4677 |
lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] = |
60606 | 4678 |
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
|
4679 |
|
56656 | 4680 |
|
60500 | 4681 |
subsection \<open>Size setup\<close> |
56656 | 4682 |
|
67332 | 4683 |
lemma size_multiset_o_map: "size_multiset g \<circ> image_mset f = size_multiset (g \<circ> f)" |
65547 | 4684 |
apply (rule ext) |
4685 |
subgoal for x by (induct x) auto |
|
4686 |
done |
|
56656 | 4687 |
|
60500 | 4688 |
setup \<open> |
69593 | 4689 |
BNF_LFP_Size.register_size_global \<^type_name>\<open>multiset\<close> \<^const_name>\<open>size_multiset\<close> |
62082 | 4690 |
@{thm size_multiset_overloaded_def} |
60606 | 4691 |
@{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single |
4692 |
size_union} |
|
67332 | 4693 |
@{thms size_multiset_o_map} |
60500 | 4694 |
\<close> |
56656 | 4695 |
|
65547 | 4696 |
subsection \<open>Lemmas about Size\<close> |
4697 |
||
4698 |
lemma size_mset_SucE: "size A = Suc n \<Longrightarrow> (\<And>a B. A = {#a#} + B \<Longrightarrow> size B = n \<Longrightarrow> P) \<Longrightarrow> P" |
|
4699 |
by (cases A) (auto simp add: ac_simps) |
|
4700 |
||
4701 |
lemma size_Suc_Diff1: "x \<in># M \<Longrightarrow> Suc (size (M - {#x#})) = size M" |
|
4702 |
using arg_cong[OF insert_DiffM, of _ _ size] by simp |
|
4703 |
||
4704 |
lemma size_Diff_singleton: "x \<in># M \<Longrightarrow> size (M - {#x#}) = size M - 1" |
|
68406 | 4705 |
by (simp flip: size_Suc_Diff1) |
65547 | 4706 |
|
4707 |
lemma size_Diff_singleton_if: "size (A - {#x#}) = (if x \<in># A then size A - 1 else size A)" |
|
4708 |
by (simp add: diff_single_trivial size_Diff_singleton) |
|
4709 |
||
4710 |
lemma size_Un_Int: "size A + size B = size (A \<union># B) + size (A \<inter># B)" |
|
4711 |
by (metis inter_subset_eq_union size_union subset_mset.diff_add union_diff_inter_eq_sup) |
|
4712 |
||
4713 |
lemma size_Un_disjoint: "A \<inter># B = {#} \<Longrightarrow> size (A \<union># B) = size A + size B" |
|
4714 |
using size_Un_Int[of A B] by simp |
|
4715 |
||
4716 |
lemma size_Diff_subset_Int: "size (M - M') = size M - size (M \<inter># M')" |
|
4717 |
by (metis diff_intersect_left_idem size_Diff_submset subset_mset.inf_le1) |
|
4718 |
||
4719 |
lemma diff_size_le_size_Diff: "size (M :: _ multiset) - size M' \<le> size (M - M')" |
|
4720 |
by (simp add: diff_le_mono2 size_Diff_subset_Int size_mset_mono) |
|
4721 |
||
4722 |
lemma size_Diff1_less: "x\<in># M \<Longrightarrow> size (M - {#x#}) < size M" |
|
4723 |
by (rule Suc_less_SucD) (simp add: size_Suc_Diff1) |
|
4724 |
||
4725 |
lemma size_Diff2_less: "x\<in># M \<Longrightarrow> y\<in># M \<Longrightarrow> size (M - {#x#} - {#y#}) < size M" |
|
4726 |
by (metis less_imp_diff_less size_Diff1_less size_Diff_subset_Int) |
|
4727 |
||
4728 |
lemma size_Diff1_le: "size (M - {#x#}) \<le> size M" |
|
4729 |
by (cases "x \<in># M") (simp_all add: size_Diff1_less less_imp_le diff_single_trivial) |
|
4730 |
||
4731 |
lemma size_psubset: "M \<subseteq># M' \<Longrightarrow> size M < size M' \<Longrightarrow> M \<subset># M'" |
|
4732 |
using less_irrefl subset_mset_def by blast |
|
4733 |
||
76700
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
blanchet
parents:
76682
diff
changeset
|
4734 |
lifting_update multiset.lifting |
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
blanchet
parents:
76682
diff
changeset
|
4735 |
lifting_forget multiset.lifting |
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
blanchet
parents:
76682
diff
changeset
|
4736 |
|
56656 | 4737 |
hide_const (open) wcount |
4738 |
||
55129
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
blanchet
parents:
54868
diff
changeset
|
4739 |
end |