| author | haftmann | 
| Sun, 24 May 2009 15:02:22 +0200 | |
| changeset 31248 | d1c65a593daf | 
| parent 30729 | 461ee3e49ad3 | 
| child 32438 | 620a5d8cfa78 | 
| permissions | -rw-r--r-- | 
| 10249 | 1 | (* Title: HOL/Library/Multiset.thy | 
| 15072 | 2 | Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker | 
| 10249 | 3 | *) | 
| 4 | ||
| 14706 | 5 | header {* Multisets *}
 | 
| 10249 | 6 | |
| 15131 | 7 | theory Multiset | 
| 30663 
0b6aff7451b2
Main is (Complex_Main) base entry point in library theories
 haftmann parents: 
30595diff
changeset | 8 | imports List Main | 
| 15131 | 9 | begin | 
| 10249 | 10 | |
| 11 | subsection {* The type of multisets *}
 | |
| 12 | ||
| 25162 | 13 | typedef 'a multiset = "{f::'a => nat. finite {x . f x > 0}}"
 | 
| 10249 | 14 | proof | 
| 11464 | 15 | show "(\<lambda>x. 0::nat) \<in> ?multiset" by simp | 
| 10249 | 16 | qed | 
| 17 | ||
| 18 | lemmas multiset_typedef [simp] = | |
| 10277 | 19 | Abs_multiset_inverse Rep_multiset_inverse Rep_multiset | 
| 20 | and [simp] = Rep_multiset_inject [symmetric] | |
| 10249 | 21 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 22 | definition Mempty :: "'a multiset"  ("{#}") where
 | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 23 |   [code del]: "{#} = Abs_multiset (\<lambda>a. 0)"
 | 
| 10249 | 24 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 25 | definition single :: "'a => 'a multiset" where | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 26 | [code del]: "single a = Abs_multiset (\<lambda>b. if b = a then 1 else 0)" | 
| 10249 | 27 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 28 | definition count :: "'a multiset => 'a => nat" where | 
| 19086 | 29 | "count = Rep_multiset" | 
| 10249 | 30 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 31 | definition MCollect :: "'a multiset => ('a => bool) => 'a multiset" where
 | 
| 19086 | 32 | "MCollect M P = Abs_multiset (\<lambda>x. if P x then Rep_multiset M x else 0)" | 
| 33 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 34 | abbreviation Melem :: "'a => 'a multiset => bool"  ("(_/ :# _)" [50, 51] 50) where
 | 
| 25610 | 35 | "a :# M == 0 < count M a" | 
| 36 | ||
| 26145 | 37 | notation (xsymbols) | 
| 38 | Melem (infix "\<in>#" 50) | |
| 10249 | 39 | |
| 40 | syntax | |
| 26033 | 41 |   "_MCollect" :: "pttrn => 'a multiset => bool => 'a multiset"    ("(1{# _ :# _./ _#})")
 | 
| 10249 | 42 | translations | 
| 26033 | 43 |   "{#x :# M. P#}" == "CONST MCollect M (\<lambda>x. P)"
 | 
| 10249 | 44 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 45 | definition set_of :: "'a multiset => 'a set" where | 
| 19086 | 46 |   "set_of M = {x. x :# M}"
 | 
| 10249 | 47 | |
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 48 | instantiation multiset :: (type) "{plus, minus, zero, size}" 
 | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 49 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 50 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 51 | definition union_def [code del]: | 
| 26145 | 52 | "M + N = Abs_multiset (\<lambda>a. Rep_multiset M a + Rep_multiset N a)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 53 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 54 | definition diff_def [code del]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 55 | "M - N = Abs_multiset (\<lambda>a. Rep_multiset M a - Rep_multiset N a)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 56 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 57 | definition Zero_multiset_def [simp]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 58 |   "0 = {#}"
 | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 59 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 60 | definition size_def: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 61 | "size M = setsum (count M) (set_of M)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 62 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 63 | instance .. | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 64 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 65 | end | 
| 10249 | 66 | |
| 19086 | 67 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21214diff
changeset | 68 | multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where | 
| 19086 | 69 | "multiset_inter A B = A - (A - B)" | 
| 15869 | 70 | |
| 26145 | 71 | text {* Multiset Enumeration *}
 | 
| 72 | syntax | |
| 26176 | 73 |   "_multiset" :: "args => 'a multiset"    ("{#(_)#}")
 | 
| 25507 | 74 | translations | 
| 75 |   "{#x, xs#}" == "{#x#} + {#xs#}"
 | |
| 76 |   "{#x#}" == "CONST single x"
 | |
| 77 | ||
| 10249 | 78 | |
| 79 | text {*
 | |
| 80 |  \medskip Preservation of the representing set @{term multiset}.
 | |
| 81 | *} | |
| 82 | ||
| 26016 | 83 | lemma const0_in_multiset: "(\<lambda>a. 0) \<in> multiset" | 
| 26178 | 84 | by (simp add: multiset_def) | 
| 10249 | 85 | |
| 26016 | 86 | lemma only1_in_multiset: "(\<lambda>b. if b = a then 1 else 0) \<in> multiset" | 
| 26178 | 87 | by (simp add: multiset_def) | 
| 10249 | 88 | |
| 26016 | 89 | lemma union_preserves_multiset: | 
| 26178 | 90 | "M \<in> multiset ==> N \<in> multiset ==> (\<lambda>a. M a + N a) \<in> multiset" | 
| 29901 | 91 | by (simp add: multiset_def) | 
| 92 | ||
| 10249 | 93 | |
| 26016 | 94 | lemma diff_preserves_multiset: | 
| 26178 | 95 | "M \<in> multiset ==> (\<lambda>a. M a - N a) \<in> multiset" | 
| 96 | apply (simp add: multiset_def) | |
| 97 | apply (rule finite_subset) | |
| 98 | apply auto | |
| 99 | done | |
| 10249 | 100 | |
| 26016 | 101 | lemma MCollect_preserves_multiset: | 
| 26178 | 102 | "M \<in> multiset ==> (\<lambda>x. if P x then M x else 0) \<in> multiset" | 
| 103 | apply (simp add: multiset_def) | |
| 104 | apply (rule finite_subset, auto) | |
| 105 | done | |
| 10249 | 106 | |
| 26016 | 107 | lemmas in_multiset = const0_in_multiset only1_in_multiset | 
| 108 | union_preserves_multiset diff_preserves_multiset MCollect_preserves_multiset | |
| 109 | ||
| 26145 | 110 | |
| 26016 | 111 | subsection {* Algebraic properties *}
 | 
| 10249 | 112 | |
| 113 | subsubsection {* Union *}
 | |
| 114 | ||
| 17161 | 115 | lemma union_empty [simp]: "M + {#} = M \<and> {#} + M = M"
 | 
| 26178 | 116 | by (simp add: union_def Mempty_def in_multiset) | 
| 10249 | 117 | |
| 17161 | 118 | lemma union_commute: "M + N = N + (M::'a multiset)" | 
| 26178 | 119 | by (simp add: union_def add_ac in_multiset) | 
| 17161 | 120 | |
| 121 | lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" | |
| 26178 | 122 | by (simp add: union_def add_ac in_multiset) | 
| 10249 | 123 | |
| 17161 | 124 | lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" | 
| 125 | proof - | |
| 26178 | 126 | have "M + (N + K) = (N + K) + M" by (rule union_commute) | 
| 127 | also have "\<dots> = N + (K + M)" by (rule union_assoc) | |
| 128 | also have "K + M = M + K" by (rule union_commute) | |
| 17161 | 129 | finally show ?thesis . | 
| 130 | qed | |
| 10249 | 131 | |
| 17161 | 132 | lemmas union_ac = union_assoc union_commute union_lcomm | 
| 10249 | 133 | |
| 14738 | 134 | instance multiset :: (type) comm_monoid_add | 
| 17200 | 135 | proof | 
| 14722 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 136 | fix a b c :: "'a multiset" | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 137 | show "(a + b) + c = a + (b + c)" by (rule union_assoc) | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 138 | show "a + b = b + a" by (rule union_commute) | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 139 | show "0 + a = a" by simp | 
| 
8e739a6eaf11
replaced apply-style proof for instance Multiset :: plus_ac0 by recommended Isar proof style
 obua parents: 
14706diff
changeset | 140 | qed | 
| 10277 | 141 | |
| 10249 | 142 | |
| 143 | subsubsection {* Difference *}
 | |
| 144 | ||
| 17161 | 145 | lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
 | 
| 26178 | 146 | by (simp add: Mempty_def diff_def in_multiset) | 
| 10249 | 147 | |
| 17161 | 148 | lemma diff_union_inverse2 [simp]: "M + {#a#} - {#a#} = M"
 | 
| 26178 | 149 | by (simp add: union_def diff_def in_multiset) | 
| 10249 | 150 | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 151 | lemma diff_cancel: "A - A = {#}"
 | 
| 26178 | 152 | by (simp add: diff_def Mempty_def) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 153 | |
| 10249 | 154 | |
| 155 | subsubsection {* Count of elements *}
 | |
| 156 | ||
| 17161 | 157 | lemma count_empty [simp]: "count {#} a = 0"
 | 
| 26178 | 158 | by (simp add: count_def Mempty_def in_multiset) | 
| 10249 | 159 | |
| 17161 | 160 | lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)"
 | 
| 26178 | 161 | by (simp add: count_def single_def in_multiset) | 
| 10249 | 162 | |
| 17161 | 163 | lemma count_union [simp]: "count (M + N) a = count M a + count N a" | 
| 26178 | 164 | by (simp add: count_def union_def in_multiset) | 
| 10249 | 165 | |
| 17161 | 166 | lemma count_diff [simp]: "count (M - N) a = count M a - count N a" | 
| 26178 | 167 | by (simp add: count_def diff_def in_multiset) | 
| 26016 | 168 | |
| 169 | lemma count_MCollect [simp]: | |
| 26178 | 170 |   "count {# x:#M. P x #} a = (if P a then count M a else 0)"
 | 
| 171 | by (simp add: count_def MCollect_def in_multiset) | |
| 10249 | 172 | |
| 173 | ||
| 174 | subsubsection {* Set of elements *}
 | |
| 175 | ||
| 17161 | 176 | lemma set_of_empty [simp]: "set_of {#} = {}"
 | 
| 26178 | 177 | by (simp add: set_of_def) | 
| 10249 | 178 | |
| 17161 | 179 | lemma set_of_single [simp]: "set_of {#b#} = {b}"
 | 
| 26178 | 180 | by (simp add: set_of_def) | 
| 10249 | 181 | |
| 17161 | 182 | lemma set_of_union [simp]: "set_of (M + N) = set_of M \<union> set_of N" | 
| 26178 | 183 | by (auto simp add: set_of_def) | 
| 10249 | 184 | |
| 17161 | 185 | lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})"
 | 
| 26818 
b4a24433154e
Instantiated rule expand_fun_eq in proof of set_of_eq_empty_iff, to avoid that
 berghofe parents: 
26567diff
changeset | 186 | by (auto simp: set_of_def Mempty_def in_multiset count_def expand_fun_eq [where f="Rep_multiset M"]) | 
| 10249 | 187 | |
| 17161 | 188 | lemma mem_set_of_iff [simp]: "(x \<in> set_of M) = (x :# M)" | 
| 26178 | 189 | by (auto simp add: set_of_def) | 
| 26016 | 190 | |
| 26033 | 191 | lemma set_of_MCollect [simp]: "set_of {# x:#M. P x #} = set_of M \<inter> {x. P x}"
 | 
| 26178 | 192 | by (auto simp add: set_of_def) | 
| 10249 | 193 | |
| 194 | ||
| 195 | subsubsection {* Size *}
 | |
| 196 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 197 | lemma size_empty [simp]: "size {#} = 0"
 | 
| 26178 | 198 | by (simp add: size_def) | 
| 10249 | 199 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 200 | lemma size_single [simp]: "size {#b#} = 1"
 | 
| 26178 | 201 | by (simp add: size_def) | 
| 10249 | 202 | |
| 17161 | 203 | lemma finite_set_of [iff]: "finite (set_of M)" | 
| 26178 | 204 | using Rep_multiset [of M] by (simp add: multiset_def set_of_def count_def) | 
| 10249 | 205 | |
| 17161 | 206 | lemma setsum_count_Int: | 
| 26178 | 207 | "finite A ==> setsum (count N) (A \<inter> set_of N) = setsum (count N) A" | 
| 208 | apply (induct rule: finite_induct) | |
| 209 | apply simp | |
| 210 | apply (simp add: Int_insert_left set_of_def) | |
| 211 | done | |
| 10249 | 212 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 213 | lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" | 
| 26178 | 214 | apply (unfold size_def) | 
| 215 | apply (subgoal_tac "count (M + N) = (\<lambda>a. count M a + count N a)") | |
| 216 | prefer 2 | |
| 217 | apply (rule ext, simp) | |
| 218 | apply (simp (no_asm_simp) add: setsum_Un_nat setsum_addf setsum_count_Int) | |
| 219 | apply (subst Int_commute) | |
| 220 | apply (simp (no_asm_simp) add: setsum_count_Int) | |
| 221 | done | |
| 10249 | 222 | |
| 17161 | 223 | lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
 | 
| 26178 | 224 | apply (unfold size_def Mempty_def count_def, auto simp: in_multiset) | 
| 225 | apply (simp add: set_of_def count_def in_multiset expand_fun_eq) | |
| 226 | done | |
| 26016 | 227 | |
| 228 | lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
 | |
| 26178 | 229 | by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) | 
| 10249 | 230 | |
| 17161 | 231 | lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \<exists>a. a :# M" | 
| 26178 | 232 | apply (unfold size_def) | 
| 233 | apply (drule setsum_SucD) | |
| 234 | apply auto | |
| 235 | done | |
| 10249 | 236 | |
| 26145 | 237 | |
| 10249 | 238 | subsubsection {* Equality of multisets *}
 | 
| 239 | ||
| 17161 | 240 | lemma multiset_eq_conv_count_eq: "(M = N) = (\<forall>a. count M a = count N a)" | 
| 26178 | 241 | by (simp add: count_def expand_fun_eq) | 
| 10249 | 242 | |
| 17161 | 243 | lemma single_not_empty [simp]: "{#a#} \<noteq> {#} \<and> {#} \<noteq> {#a#}"
 | 
| 26178 | 244 | by (simp add: single_def Mempty_def in_multiset expand_fun_eq) | 
| 10249 | 245 | |
| 17161 | 246 | lemma single_eq_single [simp]: "({#a#} = {#b#}) = (a = b)"
 | 
| 26178 | 247 | by (auto simp add: single_def in_multiset expand_fun_eq) | 
| 10249 | 248 | |
| 17161 | 249 | lemma union_eq_empty [iff]: "(M + N = {#}) = (M = {#} \<and> N = {#})"
 | 
| 26178 | 250 | by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) | 
| 10249 | 251 | |
| 17161 | 252 | lemma empty_eq_union [iff]: "({#} = M + N) = (M = {#} \<and> N = {#})"
 | 
| 26178 | 253 | by (auto simp add: union_def Mempty_def in_multiset expand_fun_eq) | 
| 10249 | 254 | |
| 17161 | 255 | lemma union_right_cancel [simp]: "(M + K = N + K) = (M = (N::'a multiset))" | 
| 26178 | 256 | by (simp add: union_def in_multiset expand_fun_eq) | 
| 10249 | 257 | |
| 17161 | 258 | lemma union_left_cancel [simp]: "(K + M = K + N) = (M = (N::'a multiset))" | 
| 26178 | 259 | by (simp add: union_def in_multiset expand_fun_eq) | 
| 10249 | 260 | |
| 17161 | 261 | lemma union_is_single: | 
| 26178 | 262 |   "(M + N = {#a#}) = (M = {#a#} \<and> N={#} \<or> M = {#} \<and> N = {#a#})"
 | 
| 263 | apply (simp add: Mempty_def single_def union_def in_multiset add_is_1 expand_fun_eq) | |
| 264 | apply blast | |
| 265 | done | |
| 10249 | 266 | |
| 17161 | 267 | lemma single_is_union: | 
| 26178 | 268 |   "({#a#} = M + N) \<longleftrightarrow> ({#a#} = M \<and> N = {#} \<or> M = {#} \<and> {#a#} = N)"
 | 
| 269 | apply (unfold Mempty_def single_def union_def) | |
| 270 | apply (simp add: add_is_1 one_is_add in_multiset expand_fun_eq) | |
| 271 | apply (blast dest: sym) | |
| 272 | done | |
| 10249 | 273 | |
| 17161 | 274 | lemma add_eq_conv_diff: | 
| 10249 | 275 |   "(M + {#a#} = N + {#b#}) =
 | 
| 15072 | 276 |    (M = N \<and> a = b \<or> M = N - {#a#} + {#b#} \<and> N = M - {#b#} + {#a#})"
 | 
| 26178 | 277 | using [[simproc del: neq]] | 
| 278 | apply (unfold single_def union_def diff_def) | |
| 279 | apply (simp (no_asm) add: in_multiset expand_fun_eq) | |
| 280 | apply (rule conjI, force, safe, simp_all) | |
| 281 | apply (simp add: eq_sym_conv) | |
| 282 | done | |
| 10249 | 283 | |
| 15869 | 284 | declare Rep_multiset_inject [symmetric, simp del] | 
| 285 | ||
| 23611 | 286 | instance multiset :: (type) cancel_ab_semigroup_add | 
| 287 | proof | |
| 288 | fix a b c :: "'a multiset" | |
| 289 | show "a + b = a + c \<Longrightarrow> b = c" by simp | |
| 290 | qed | |
| 15869 | 291 | |
| 25610 | 292 | lemma insert_DiffM: | 
| 293 |   "x \<in># M \<Longrightarrow> {#x#} + (M - {#x#}) = M"
 | |
| 26178 | 294 | by (clarsimp simp: multiset_eq_conv_count_eq) | 
| 25610 | 295 | |
| 296 | lemma insert_DiffM2[simp]: | |
| 297 |   "x \<in># M \<Longrightarrow> M - {#x#} + {#x#} = M"
 | |
| 26178 | 298 | by (clarsimp simp: multiset_eq_conv_count_eq) | 
| 25610 | 299 | |
| 300 | lemma multi_union_self_other_eq: | |
| 301 | "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y" | |
| 26178 | 302 | by (induct A arbitrary: X Y) auto | 
| 25610 | 303 | |
| 304 | lemma multi_self_add_other_not_self[simp]: "(A = A + {#x#}) = False"
 | |
| 26178 | 305 | by (metis single_not_empty union_empty union_left_cancel) | 
| 25610 | 306 | |
| 307 | lemma insert_noteq_member: | |
| 308 |   assumes BC: "B + {#b#} = C + {#c#}"
 | |
| 309 | and bnotc: "b \<noteq> c" | |
| 310 | shows "c \<in># B" | |
| 311 | proof - | |
| 312 |   have "c \<in># C + {#c#}" by simp
 | |
| 313 |   have nc: "\<not> c \<in># {#b#}" using bnotc by simp
 | |
| 26145 | 314 |   then have "c \<in># B + {#b#}" using BC by simp
 | 
| 315 | then show "c \<in># B" using nc by simp | |
| 25610 | 316 | qed | 
| 317 | ||
| 318 | ||
| 26016 | 319 | lemma add_eq_conv_ex: | 
| 320 |   "(M + {#a#} = N + {#b#}) =
 | |
| 321 |     (M = N \<and> a = b \<or> (\<exists>K. M = K + {#b#} \<and> N = K + {#a#}))"
 | |
| 26178 | 322 | by (auto simp add: add_eq_conv_diff) | 
| 26016 | 323 | |
| 324 | ||
| 325 | lemma empty_multiset_count: | |
| 326 |   "(\<forall>x. count A x = 0) = (A = {#})"
 | |
| 26178 | 327 | by (metis count_empty multiset_eq_conv_count_eq) | 
| 26016 | 328 | |
| 329 | ||
| 15869 | 330 | subsubsection {* Intersection *}
 | 
| 331 | ||
| 332 | lemma multiset_inter_count: | |
| 26178 | 333 | "count (A #\<inter> B) x = min (count A x) (count B x)" | 
| 334 | by (simp add: multiset_inter_def min_def) | |
| 15869 | 335 | |
| 336 | lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A" | |
| 26178 | 337 | by (simp add: multiset_eq_conv_count_eq multiset_inter_count | 
| 21214 
a91bab12b2bd
adjusted two lemma names due to name change in interpretation
 haftmann parents: 
20770diff
changeset | 338 | min_max.inf_commute) | 
| 15869 | 339 | |
| 340 | lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C" | |
| 26178 | 341 | by (simp add: multiset_eq_conv_count_eq multiset_inter_count | 
| 21214 
a91bab12b2bd
adjusted two lemma names due to name change in interpretation
 haftmann parents: 
20770diff
changeset | 342 | min_max.inf_assoc) | 
| 15869 | 343 | |
| 344 | lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)" | |
| 26178 | 345 | by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def) | 
| 15869 | 346 | |
| 17161 | 347 | lemmas multiset_inter_ac = | 
| 348 | multiset_inter_commute | |
| 349 | multiset_inter_assoc | |
| 350 | multiset_inter_left_commute | |
| 15869 | 351 | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 352 | lemma multiset_inter_single: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
 | 
| 26178 | 353 | by (simp add: multiset_eq_conv_count_eq multiset_inter_count) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 354 | |
| 15869 | 355 | lemma multiset_union_diff_commute: "B #\<inter> C = {#} \<Longrightarrow> A + B - C = A - C + B"
 | 
| 26178 | 356 | apply (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def | 
| 17161 | 357 | split: split_if_asm) | 
| 26178 | 358 | apply clarsimp | 
| 359 | apply (erule_tac x = a in allE) | |
| 360 | apply auto | |
| 361 | done | |
| 15869 | 362 | |
| 10249 | 363 | |
| 26016 | 364 | subsubsection {* Comprehension (filter) *}
 | 
| 365 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 366 | lemma MCollect_empty [simp]: "MCollect {#} P = {#}"
 | 
| 26178 | 367 | by (simp add: MCollect_def Mempty_def Abs_multiset_inject | 
| 26145 | 368 | in_multiset expand_fun_eq) | 
| 26016 | 369 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 370 | lemma MCollect_single [simp]: | 
| 26178 | 371 |   "MCollect {#x#} P = (if P x then {#x#} else {#})"
 | 
| 372 | by (simp add: MCollect_def Mempty_def single_def Abs_multiset_inject | |
| 26145 | 373 | in_multiset expand_fun_eq) | 
| 26016 | 374 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 375 | lemma MCollect_union [simp]: | 
| 26016 | 376 | "MCollect (M+N) f = MCollect M f + MCollect N f" | 
| 26178 | 377 | by (simp add: MCollect_def union_def Abs_multiset_inject | 
| 26145 | 378 | in_multiset expand_fun_eq) | 
| 26016 | 379 | |
| 380 | ||
| 381 | subsection {* Induction and case splits *}
 | |
| 10249 | 382 | |
| 383 | lemma setsum_decr: | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11655diff
changeset | 384 | "finite F ==> (0::nat) < f a ==> | 
| 15072 | 385 | setsum (f (a := f a - 1)) F = (if a\<in>F then setsum f F - 1 else setsum f F)" | 
| 26178 | 386 | apply (induct rule: finite_induct) | 
| 387 | apply auto | |
| 388 | apply (drule_tac a = a in mk_disjoint_insert, auto) | |
| 389 | done | |
| 10249 | 390 | |
| 10313 | 391 | lemma rep_multiset_induct_aux: | 
| 26178 | 392 | assumes 1: "P (\<lambda>a. (0::nat))" | 
| 393 | and 2: "!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))" | |
| 394 | shows "\<forall>f. f \<in> multiset --> setsum f {x. f x \<noteq> 0} = n --> P f"
 | |
| 395 | apply (unfold multiset_def) | |
| 396 | apply (induct_tac n, simp, clarify) | |
| 397 | apply (subgoal_tac "f = (\<lambda>a.0)") | |
| 398 | apply simp | |
| 399 | apply (rule 1) | |
| 400 | apply (rule ext, force, clarify) | |
| 401 | apply (frule setsum_SucD, clarify) | |
| 402 | apply (rename_tac a) | |
| 403 | apply (subgoal_tac "finite {x. (f (a := f a - 1)) x > 0}")
 | |
| 404 | prefer 2 | |
| 405 | apply (rule finite_subset) | |
| 406 | prefer 2 | |
| 407 | apply assumption | |
| 408 | apply simp | |
| 409 | apply blast | |
| 410 | apply (subgoal_tac "f = (f (a := f a - 1))(a := (f (a := f a - 1)) a + 1)") | |
| 411 | prefer 2 | |
| 412 | apply (rule ext) | |
| 413 | apply (simp (no_asm_simp)) | |
| 414 | apply (erule ssubst, rule 2 [unfolded multiset_def], blast) | |
| 415 | apply (erule allE, erule impE, erule_tac [2] mp, blast) | |
| 416 | apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply One_nat_def) | |
| 417 | apply (subgoal_tac "{x. x \<noteq> a --> f x \<noteq> 0} = {x. f x \<noteq> 0}")
 | |
| 418 | prefer 2 | |
| 419 | apply blast | |
| 420 | apply (subgoal_tac "{x. x \<noteq> a \<and> f x \<noteq> 0} = {x. f x \<noteq> 0} - {a}")
 | |
| 421 | prefer 2 | |
| 422 | apply blast | |
| 423 | apply (simp add: le_imp_diff_is_add setsum_diff1_nat cong: conj_cong) | |
| 424 | done | |
| 10249 | 425 | |
| 10313 | 426 | theorem rep_multiset_induct: | 
| 11464 | 427 | "f \<in> multiset ==> P (\<lambda>a. 0) ==> | 
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11655diff
changeset | 428 | (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))) ==> P f" | 
| 26178 | 429 | using rep_multiset_induct_aux by blast | 
| 10249 | 430 | |
| 18258 | 431 | theorem multiset_induct [case_names empty add, induct type: multiset]: | 
| 26178 | 432 | assumes empty: "P {#}"
 | 
| 433 |   and add: "!!M x. P M ==> P (M + {#x#})"
 | |
| 434 | shows "P M" | |
| 10249 | 435 | proof - | 
| 436 | note defns = union_def single_def Mempty_def | |
| 437 | show ?thesis | |
| 438 | apply (rule Rep_multiset_inverse [THEN subst]) | |
| 10313 | 439 | apply (rule Rep_multiset [THEN rep_multiset_induct]) | 
| 18258 | 440 | apply (rule empty [unfolded defns]) | 
| 15072 | 441 | apply (subgoal_tac "f(b := f b + 1) = (\<lambda>a. f a + (if a=b then 1 else 0))") | 
| 10249 | 442 | prefer 2 | 
| 443 | apply (simp add: expand_fun_eq) | |
| 444 | apply (erule ssubst) | |
| 17200 | 445 | apply (erule Abs_multiset_inverse [THEN subst]) | 
| 26016 | 446 | apply (drule add [unfolded defns, simplified]) | 
| 447 | apply(simp add:in_multiset) | |
| 10249 | 448 | done | 
| 449 | qed | |
| 450 | ||
| 25610 | 451 | lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = A + {#a#}"
 | 
| 26178 | 452 | by (induct M) auto | 
| 25610 | 453 | |
| 454 | lemma multiset_cases [cases type, case_names empty add]: | |
| 26178 | 455 | assumes em:  "M = {#} \<Longrightarrow> P"
 | 
| 456 | assumes add: "\<And>N x. M = N + {#x#} \<Longrightarrow> P"
 | |
| 457 | shows "P" | |
| 25610 | 458 | proof (cases "M = {#}")
 | 
| 26145 | 459 |   assume "M = {#}" then show ?thesis using em by simp
 | 
| 25610 | 460 | next | 
| 461 |   assume "M \<noteq> {#}"
 | |
| 462 |   then obtain M' m where "M = M' + {#m#}" 
 | |
| 463 | by (blast dest: multi_nonempty_split) | |
| 26145 | 464 | then show ?thesis using add by simp | 
| 25610 | 465 | qed | 
| 466 | ||
| 467 | lemma multi_member_split: "x \<in># M \<Longrightarrow> \<exists>A. M = A + {#x#}"
 | |
| 26178 | 468 | apply (cases M) | 
| 469 | apply simp | |
| 470 | apply (rule_tac x="M - {#x#}" in exI, simp)
 | |
| 471 | done | |
| 25610 | 472 | |
| 26033 | 473 | lemma multiset_partition: "M = {# x:#M. P x #} + {# x:#M. \<not> P x #}"
 | 
| 26178 | 474 | apply (subst multiset_eq_conv_count_eq) | 
| 475 | apply auto | |
| 476 | done | |
| 10249 | 477 | |
| 15869 | 478 | declare multiset_typedef [simp del] | 
| 10249 | 479 | |
| 25610 | 480 | lemma multi_drop_mem_not_eq: "c \<in># B \<Longrightarrow> B - {#c#} \<noteq> B"
 | 
| 26178 | 481 | by (cases "B = {#}") (auto dest: multi_member_split)
 | 
| 26145 | 482 | |
| 17161 | 483 | |
| 26016 | 484 | subsection {* Orderings *}
 | 
| 10249 | 485 | |
| 486 | subsubsection {* Well-foundedness *}
 | |
| 487 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 488 | definition mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
 | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 489 |   [code del]: "mult1 r = {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
 | 
| 23751 | 490 | (\<forall>b. b :# K --> (b, a) \<in> r)}" | 
| 10249 | 491 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 492 | definition mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
 | 
| 23751 | 493 | "mult r = (mult1 r)\<^sup>+" | 
| 10249 | 494 | |
| 23751 | 495 | lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
 | 
| 26178 | 496 | by (simp add: mult1_def) | 
| 10249 | 497 | |
| 23751 | 498 | lemma less_add: "(N, M0 + {#a#}) \<in> mult1 r ==>
 | 
| 499 |     (\<exists>M. (M, M0) \<in> mult1 r \<and> N = M + {#a#}) \<or>
 | |
| 500 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K)" | |
| 19582 | 501 | (is "_ \<Longrightarrow> ?case1 (mult1 r) \<or> ?case2") | 
| 10249 | 502 | proof (unfold mult1_def) | 
| 23751 | 503 | let ?r = "\<lambda>K a. \<forall>b. b :# K --> (b, a) \<in> r" | 
| 11464 | 504 |   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
 | 
| 23751 | 505 |   let ?case1 = "?case1 {(N, M). ?R N M}"
 | 
| 10249 | 506 | |
| 23751 | 507 |   assume "(N, M0 + {#a#}) \<in> {(N, M). ?R N M}"
 | 
| 18258 | 508 | then have "\<exists>a' M0' K. | 
| 11464 | 509 |       M0 + {#a#} = M0' + {#a'#} \<and> N = M0' + K \<and> ?r K a'" by simp
 | 
| 18258 | 510 | then show "?case1 \<or> ?case2" | 
| 10249 | 511 | proof (elim exE conjE) | 
| 512 | fix a' M0' K | |
| 513 | assume N: "N = M0' + K" and r: "?r K a'" | |
| 514 |     assume "M0 + {#a#} = M0' + {#a'#}"
 | |
| 18258 | 515 | then have "M0 = M0' \<and> a = a' \<or> | 
| 11464 | 516 |         (\<exists>K'. M0 = K' + {#a'#} \<and> M0' = K' + {#a#})"
 | 
| 10249 | 517 | by (simp only: add_eq_conv_ex) | 
| 18258 | 518 | then show ?thesis | 
| 10249 | 519 | proof (elim disjE conjE exE) | 
| 520 | assume "M0 = M0'" "a = a'" | |
| 11464 | 521 | with N r have "?r K a \<and> N = M0 + K" by simp | 
| 18258 | 522 | then have ?case2 .. then show ?thesis .. | 
| 10249 | 523 | next | 
| 524 | fix K' | |
| 525 |       assume "M0' = K' + {#a#}"
 | |
| 526 |       with N have n: "N = K' + K + {#a#}" by (simp add: union_ac)
 | |
| 527 | ||
| 528 |       assume "M0 = K' + {#a'#}"
 | |
| 529 | with r have "?R (K' + K) M0" by blast | |
| 18258 | 530 | with n have ?case1 by simp then show ?thesis .. | 
| 10249 | 531 | qed | 
| 532 | qed | |
| 533 | qed | |
| 534 | ||
| 23751 | 535 | lemma all_accessible: "wf r ==> \<forall>M. M \<in> acc (mult1 r)" | 
| 10249 | 536 | proof | 
| 537 | let ?R = "mult1 r" | |
| 538 | let ?W = "acc ?R" | |
| 539 |   {
 | |
| 540 | fix M M0 a | |
| 23751 | 541 | assume M0: "M0 \<in> ?W" | 
| 542 |       and wf_hyp: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | |
| 543 |       and acc_hyp: "\<forall>M. (M, M0) \<in> ?R --> M + {#a#} \<in> ?W"
 | |
| 544 |     have "M0 + {#a#} \<in> ?W"
 | |
| 545 |     proof (rule accI [of "M0 + {#a#}"])
 | |
| 10249 | 546 | fix N | 
| 23751 | 547 |       assume "(N, M0 + {#a#}) \<in> ?R"
 | 
| 548 |       then have "((\<exists>M. (M, M0) \<in> ?R \<and> N = M + {#a#}) \<or>
 | |
| 549 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K))" | |
| 10249 | 550 | by (rule less_add) | 
| 23751 | 551 | then show "N \<in> ?W" | 
| 10249 | 552 | proof (elim exE disjE conjE) | 
| 23751 | 553 |         fix M assume "(M, M0) \<in> ?R" and N: "N = M + {#a#}"
 | 
| 554 |         from acc_hyp have "(M, M0) \<in> ?R --> M + {#a#} \<in> ?W" ..
 | |
| 555 |         from this and `(M, M0) \<in> ?R` have "M + {#a#} \<in> ?W" ..
 | |
| 556 | then show "N \<in> ?W" by (simp only: N) | |
| 10249 | 557 | next | 
| 558 | fix K | |
| 559 | assume N: "N = M0 + K" | |
| 23751 | 560 | assume "\<forall>b. b :# K --> (b, a) \<in> r" | 
| 561 | then have "M0 + K \<in> ?W" | |
| 10249 | 562 | proof (induct K) | 
| 18730 | 563 | case empty | 
| 23751 | 564 |           from M0 show "M0 + {#} \<in> ?W" by simp
 | 
| 18730 | 565 | next | 
| 566 | case (add K x) | |
| 23751 | 567 | from add.prems have "(x, a) \<in> r" by simp | 
| 568 |           with wf_hyp have "\<forall>M \<in> ?W. M + {#x#} \<in> ?W" by blast
 | |
| 569 | moreover from add have "M0 + K \<in> ?W" by simp | |
| 570 |           ultimately have "(M0 + K) + {#x#} \<in> ?W" ..
 | |
| 571 |           then show "M0 + (K + {#x#}) \<in> ?W" by (simp only: union_assoc)
 | |
| 10249 | 572 | qed | 
| 23751 | 573 | then show "N \<in> ?W" by (simp only: N) | 
| 10249 | 574 | qed | 
| 575 | qed | |
| 576 | } note tedious_reasoning = this | |
| 577 | ||
| 23751 | 578 | assume wf: "wf r" | 
| 10249 | 579 | fix M | 
| 23751 | 580 | show "M \<in> ?W" | 
| 10249 | 581 | proof (induct M) | 
| 23751 | 582 |     show "{#} \<in> ?W"
 | 
| 10249 | 583 | proof (rule accI) | 
| 23751 | 584 |       fix b assume "(b, {#}) \<in> ?R"
 | 
| 585 | with not_less_empty show "b \<in> ?W" by contradiction | |
| 10249 | 586 | qed | 
| 587 | ||
| 23751 | 588 | fix M a assume "M \<in> ?W" | 
| 589 |     from wf have "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | |
| 10249 | 590 | proof induct | 
| 591 | fix a | |
| 23751 | 592 |       assume r: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | 
| 593 |       show "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | |
| 10249 | 594 | proof | 
| 23751 | 595 | fix M assume "M \<in> ?W" | 
| 596 |         then show "M + {#a#} \<in> ?W"
 | |
| 23373 | 597 | by (rule acc_induct) (rule tedious_reasoning [OF _ r]) | 
| 10249 | 598 | qed | 
| 599 | qed | |
| 23751 | 600 |     from this and `M \<in> ?W` show "M + {#a#} \<in> ?W" ..
 | 
| 10249 | 601 | qed | 
| 602 | qed | |
| 603 | ||
| 23751 | 604 | theorem wf_mult1: "wf r ==> wf (mult1 r)" | 
| 26178 | 605 | by (rule acc_wfI) (rule all_accessible) | 
| 10249 | 606 | |
| 23751 | 607 | theorem wf_mult: "wf r ==> wf (mult r)" | 
| 26178 | 608 | unfolding mult_def by (rule wf_trancl) (rule wf_mult1) | 
| 10249 | 609 | |
| 610 | ||
| 611 | subsubsection {* Closure-free presentation *}
 | |
| 612 | ||
| 613 | (*Badly needed: a linear arithmetic procedure for multisets*) | |
| 614 | ||
| 615 | lemma diff_union_single_conv: "a :# J ==> I + J - {#a#} = I + (J - {#a#})"
 | |
| 26178 | 616 | by (simp add: multiset_eq_conv_count_eq) | 
| 10249 | 617 | |
| 618 | text {* One direction. *}
 | |
| 619 | ||
| 620 | lemma mult_implies_one_step: | |
| 23751 | 621 | "trans r ==> (M, N) \<in> mult r ==> | 
| 11464 | 622 |     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
 | 
| 23751 | 623 | (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r)" | 
| 26178 | 624 | apply (unfold mult_def mult1_def set_of_def) | 
| 625 | apply (erule converse_trancl_induct, clarify) | |
| 626 | apply (rule_tac x = M0 in exI, simp, clarify) | |
| 627 | apply (case_tac "a :# K") | |
| 628 | apply (rule_tac x = I in exI) | |
| 629 | apply (simp (no_asm)) | |
| 630 |  apply (rule_tac x = "(K - {#a#}) + Ka" in exI)
 | |
| 631 | apply (simp (no_asm_simp) add: union_assoc [symmetric]) | |
| 632 |  apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong)
 | |
| 633 | apply (simp add: diff_union_single_conv) | |
| 634 | apply (simp (no_asm_use) add: trans_def) | |
| 635 | apply blast | |
| 636 | apply (subgoal_tac "a :# I") | |
| 637 |  apply (rule_tac x = "I - {#a#}" in exI)
 | |
| 638 |  apply (rule_tac x = "J + {#a#}" in exI)
 | |
| 639 | apply (rule_tac x = "K + Ka" in exI) | |
| 640 | apply (rule conjI) | |
| 641 | apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) | |
| 642 | apply (rule conjI) | |
| 643 |   apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong, simp)
 | |
| 644 | apply (simp add: multiset_eq_conv_count_eq split: nat_diff_split) | |
| 645 | apply (simp (no_asm_use) add: trans_def) | |
| 646 | apply blast | |
| 647 | apply (subgoal_tac "a :# (M0 + {#a#})")
 | |
| 648 | apply simp | |
| 649 | apply (simp (no_asm)) | |
| 650 | done | |
| 10249 | 651 | |
| 652 | lemma elem_imp_eq_diff_union: "a :# M ==> M = M - {#a#} + {#a#}"
 | |
| 26178 | 653 | by (simp add: multiset_eq_conv_count_eq) | 
| 10249 | 654 | |
| 11464 | 655 | lemma size_eq_Suc_imp_eq_union: "size M = Suc n ==> \<exists>a N. M = N + {#a#}"
 | 
| 26178 | 656 | apply (erule size_eq_Suc_imp_elem [THEN exE]) | 
| 657 | apply (drule elem_imp_eq_diff_union, auto) | |
| 658 | done | |
| 10249 | 659 | |
| 660 | lemma one_step_implies_mult_aux: | |
| 23751 | 661 | "trans r ==> | 
| 662 |     \<forall>I J K. (size J = n \<and> J \<noteq> {#} \<and> (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r))
 | |
| 663 | --> (I + K, I + J) \<in> mult r" | |
| 26178 | 664 | apply (induct_tac n, auto) | 
| 665 | apply (frule size_eq_Suc_imp_eq_union, clarify) | |
| 666 | apply (rename_tac "J'", simp) | |
| 667 | apply (erule notE, auto) | |
| 668 | apply (case_tac "J' = {#}")
 | |
| 669 | apply (simp add: mult_def) | |
| 670 | apply (rule r_into_trancl) | |
| 671 | apply (simp add: mult1_def set_of_def, blast) | |
| 672 | txt {* Now we know @{term "J' \<noteq> {#}"}. *}
 | |
| 673 | apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition) | |
| 674 | apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp) | |
| 675 | apply (erule ssubst) | |
| 676 | apply (simp add: Ball_def, auto) | |
| 677 | apply (subgoal_tac | |
| 678 |   "((I + {# x :# K. (x, a) \<in> r #}) + {# x :# K. (x, a) \<notin> r #},
 | |
| 679 |     (I + {# x :# K. (x, a) \<in> r #}) + J') \<in> mult r")
 | |
| 680 | prefer 2 | |
| 681 | apply force | |
| 682 | apply (simp (no_asm_use) add: union_assoc [symmetric] mult_def) | |
| 683 | apply (erule trancl_trans) | |
| 684 | apply (rule r_into_trancl) | |
| 685 | apply (simp add: mult1_def set_of_def) | |
| 686 | apply (rule_tac x = a in exI) | |
| 687 | apply (rule_tac x = "I + J'" in exI) | |
| 688 | apply (simp add: union_ac) | |
| 689 | done | |
| 10249 | 690 | |
| 17161 | 691 | lemma one_step_implies_mult: | 
| 23751 | 692 |   "trans r ==> J \<noteq> {#} ==> \<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r
 | 
| 693 | ==> (I + K, I + J) \<in> mult r" | |
| 26178 | 694 | using one_step_implies_mult_aux by blast | 
| 10249 | 695 | |
| 696 | ||
| 697 | subsubsection {* Partial-order properties *}
 | |
| 698 | ||
| 26567 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 699 | instantiation multiset :: (order) order | 
| 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 700 | begin | 
| 10249 | 701 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 702 | definition less_multiset_def [code del]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 703 |   "M' < M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
 | 
| 26567 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 704 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 705 | definition le_multiset_def [code del]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 706 | "M' <= M \<longleftrightarrow> M' = M \<or> M' < (M::'a multiset)" | 
| 10249 | 707 | |
| 23751 | 708 | lemma trans_base_order: "trans {(x', x). x' < (x::'a::order)}"
 | 
| 26178 | 709 | unfolding trans_def by (blast intro: order_less_trans) | 
| 10249 | 710 | |
| 711 | text {*
 | |
| 712 | \medskip Irreflexivity. | |
| 713 | *} | |
| 714 | ||
| 715 | lemma mult_irrefl_aux: | |
| 26178 | 716 |   "finite A ==> (\<forall>x \<in> A. \<exists>y \<in> A. x < (y::'a::order)) \<Longrightarrow> A = {}"
 | 
| 717 | by (induct rule: finite_induct) (auto intro: order_less_trans) | |
| 10249 | 718 | |
| 17161 | 719 | lemma mult_less_not_refl: "\<not> M < (M::'a::order multiset)" | 
| 26178 | 720 | apply (unfold less_multiset_def, auto) | 
| 721 | apply (drule trans_base_order [THEN mult_implies_one_step], auto) | |
| 722 | apply (drule finite_set_of [THEN mult_irrefl_aux [rule_format (no_asm)]]) | |
| 723 | apply (simp add: set_of_eq_empty_iff) | |
| 724 | done | |
| 10249 | 725 | |
| 726 | lemma mult_less_irrefl [elim!]: "M < (M::'a::order multiset) ==> R" | |
| 26178 | 727 | using insert mult_less_not_refl by fast | 
| 10249 | 728 | |
| 729 | ||
| 730 | text {* Transitivity. *}
 | |
| 731 | ||
| 732 | theorem mult_less_trans: "K < M ==> M < N ==> K < (N::'a::order multiset)" | |
| 26178 | 733 | unfolding less_multiset_def mult_def by (blast intro: trancl_trans) | 
| 10249 | 734 | |
| 735 | text {* Asymmetry. *}
 | |
| 736 | ||
| 11464 | 737 | theorem mult_less_not_sym: "M < N ==> \<not> N < (M::'a::order multiset)" | 
| 26178 | 738 | apply auto | 
| 739 | apply (rule mult_less_not_refl [THEN notE]) | |
| 740 | apply (erule mult_less_trans, assumption) | |
| 741 | done | |
| 10249 | 742 | |
| 743 | theorem mult_less_asym: | |
| 26178 | 744 | "M < N ==> (\<not> P ==> N < (M::'a::order multiset)) ==> P" | 
| 745 | using mult_less_not_sym by blast | |
| 10249 | 746 | |
| 747 | theorem mult_le_refl [iff]: "M <= (M::'a::order multiset)" | |
| 26178 | 748 | unfolding le_multiset_def by auto | 
| 10249 | 749 | |
| 750 | text {* Anti-symmetry. *}
 | |
| 751 | ||
| 752 | theorem mult_le_antisym: | |
| 26178 | 753 | "M <= N ==> N <= M ==> M = (N::'a::order multiset)" | 
| 754 | unfolding le_multiset_def by (blast dest: mult_less_not_sym) | |
| 10249 | 755 | |
| 756 | text {* Transitivity. *}
 | |
| 757 | ||
| 758 | theorem mult_le_trans: | |
| 26178 | 759 | "K <= M ==> M <= N ==> K <= (N::'a::order multiset)" | 
| 760 | unfolding le_multiset_def by (blast intro: mult_less_trans) | |
| 10249 | 761 | |
| 11655 | 762 | theorem mult_less_le: "(M < N) = (M <= N \<and> M \<noteq> (N::'a::order multiset))" | 
| 26178 | 763 | unfolding le_multiset_def by auto | 
| 10249 | 764 | |
| 27682 | 765 | instance proof | 
| 766 | qed (auto simp add: mult_less_le dest: mult_le_antisym elim: mult_le_trans) | |
| 10277 | 767 | |
| 26567 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 768 | end | 
| 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 769 | |
| 10249 | 770 | |
| 771 | subsubsection {* Monotonicity of multiset union *}
 | |
| 772 | ||
| 17161 | 773 | lemma mult1_union: | 
| 26178 | 774 | "(B, D) \<in> mult1 r ==> trans r ==> (C + B, C + D) \<in> mult1 r" | 
| 775 | apply (unfold mult1_def) | |
| 776 | apply auto | |
| 777 | apply (rule_tac x = a in exI) | |
| 778 | apply (rule_tac x = "C + M0" in exI) | |
| 779 | apply (simp add: union_assoc) | |
| 780 | done | |
| 10249 | 781 | |
| 782 | lemma union_less_mono2: "B < D ==> C + B < C + (D::'a::order multiset)" | |
| 26178 | 783 | apply (unfold less_multiset_def mult_def) | 
| 784 | apply (erule trancl_induct) | |
| 785 | apply (blast intro: mult1_union transI order_less_trans r_into_trancl) | |
| 786 | apply (blast intro: mult1_union transI order_less_trans r_into_trancl trancl_trans) | |
| 787 | done | |
| 10249 | 788 | |
| 789 | lemma union_less_mono1: "B < D ==> B + C < D + (C::'a::order multiset)" | |
| 26178 | 790 | apply (subst union_commute [of B C]) | 
| 791 | apply (subst union_commute [of D C]) | |
| 792 | apply (erule union_less_mono2) | |
| 793 | done | |
| 10249 | 794 | |
| 17161 | 795 | lemma union_less_mono: | 
| 26178 | 796 | "A < C ==> B < D ==> A + B < C + (D::'a::order multiset)" | 
| 797 | by (blast intro!: union_less_mono1 union_less_mono2 mult_less_trans) | |
| 10249 | 798 | |
| 17161 | 799 | lemma union_le_mono: | 
| 26178 | 800 | "A <= C ==> B <= D ==> A + B <= C + (D::'a::order multiset)" | 
| 801 | unfolding le_multiset_def | |
| 802 | by (blast intro: union_less_mono union_less_mono1 union_less_mono2) | |
| 10249 | 803 | |
| 17161 | 804 | lemma empty_leI [iff]: "{#} <= (M::'a::order multiset)"
 | 
| 26178 | 805 | apply (unfold le_multiset_def less_multiset_def) | 
| 806 | apply (case_tac "M = {#}")
 | |
| 807 | prefer 2 | |
| 808 |  apply (subgoal_tac "({#} + {#}, {#} + M) \<in> mult (Collect (split op <))")
 | |
| 809 | prefer 2 | |
| 810 | apply (rule one_step_implies_mult) | |
| 811 | apply (simp only: trans_def) | |
| 812 | apply auto | |
| 813 | done | |
| 10249 | 814 | |
| 17161 | 815 | lemma union_upper1: "A <= A + (B::'a::order multiset)" | 
| 15072 | 816 | proof - | 
| 17200 | 817 |   have "A + {#} <= A + B" by (blast intro: union_le_mono)
 | 
| 18258 | 818 | then show ?thesis by simp | 
| 15072 | 819 | qed | 
| 820 | ||
| 17161 | 821 | lemma union_upper2: "B <= A + (B::'a::order multiset)" | 
| 26178 | 822 | by (subst union_commute) (rule union_upper1) | 
| 15072 | 823 | |
| 23611 | 824 | instance multiset :: (order) pordered_ab_semigroup_add | 
| 26178 | 825 | apply intro_classes | 
| 826 | apply (erule union_le_mono[OF mult_le_refl]) | |
| 827 | done | |
| 26145 | 828 | |
| 15072 | 829 | |
| 17200 | 830 | subsection {* Link with lists *}
 | 
| 15072 | 831 | |
| 26016 | 832 | primrec multiset_of :: "'a list \<Rightarrow> 'a multiset" where | 
| 26145 | 833 |   "multiset_of [] = {#}" |
 | 
| 834 |   "multiset_of (a # x) = multiset_of x + {# a #}"
 | |
| 15072 | 835 | |
| 836 | lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])"
 | |
| 26178 | 837 | by (induct x) auto | 
| 15072 | 838 | |
| 839 | lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])"
 | |
| 26178 | 840 | by (induct x) auto | 
| 15072 | 841 | |
| 842 | lemma set_of_multiset_of[simp]: "set_of(multiset_of x) = set x" | |
| 26178 | 843 | by (induct x) auto | 
| 15867 | 844 | |
| 845 | lemma mem_set_multiset_eq: "x \<in> set xs = (x :# multiset_of xs)" | |
| 26178 | 846 | by (induct xs) auto | 
| 15072 | 847 | |
| 18258 | 848 | lemma multiset_of_append [simp]: | 
| 26178 | 849 | "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" | 
| 850 | by (induct xs arbitrary: ys) (auto simp: union_ac) | |
| 18730 | 851 | |
| 15072 | 852 | lemma surj_multiset_of: "surj multiset_of" | 
| 26178 | 853 | apply (unfold surj_def) | 
| 854 | apply (rule allI) | |
| 855 | apply (rule_tac M = y in multiset_induct) | |
| 856 | apply auto | |
| 857 | apply (rule_tac x = "x # xa" in exI) | |
| 858 | apply auto | |
| 859 | done | |
| 10249 | 860 | |
| 25162 | 861 | lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}"
 | 
| 26178 | 862 | by (induct x) auto | 
| 15072 | 863 | |
| 17200 | 864 | lemma distinct_count_atmost_1: | 
| 26178 | 865 | "distinct x = (! a. count (multiset_of x) a = (if a \<in> set x then 1 else 0))" | 
| 866 | apply (induct x, simp, rule iffI, simp_all) | |
| 867 | apply (rule conjI) | |
| 868 | apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) | |
| 869 | apply (erule_tac x = a in allE, simp, clarify) | |
| 870 | apply (erule_tac x = aa in allE, simp) | |
| 871 | done | |
| 15072 | 872 | |
| 17200 | 873 | lemma multiset_of_eq_setD: | 
| 15867 | 874 | "multiset_of xs = multiset_of ys \<Longrightarrow> set xs = set ys" | 
| 26178 | 875 | by (rule) (auto simp add:multiset_eq_conv_count_eq set_count_greater_0) | 
| 15867 | 876 | |
| 17200 | 877 | lemma set_eq_iff_multiset_of_eq_distinct: | 
| 26145 | 878 | "distinct x \<Longrightarrow> distinct y \<Longrightarrow> | 
| 879 | (set x = set y) = (multiset_of x = multiset_of y)" | |
| 26178 | 880 | by (auto simp: multiset_eq_conv_count_eq distinct_count_atmost_1) | 
| 15072 | 881 | |
| 17200 | 882 | lemma set_eq_iff_multiset_of_remdups_eq: | 
| 15072 | 883 | "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" | 
| 26178 | 884 | apply (rule iffI) | 
| 885 | apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) | |
| 886 | apply (drule distinct_remdups [THEN distinct_remdups | |
| 26145 | 887 | [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]]) | 
| 26178 | 888 | apply simp | 
| 889 | done | |
| 10249 | 890 | |
| 18258 | 891 | lemma multiset_of_compl_union [simp]: | 
| 26178 | 892 | "multiset_of [x\<leftarrow>xs. P x] + multiset_of [x\<leftarrow>xs. \<not>P x] = multiset_of xs" | 
| 893 | by (induct xs) (auto simp: union_ac) | |
| 15072 | 894 | |
| 17200 | 895 | lemma count_filter: | 
| 26178 | 896 | "count (multiset_of xs) x = length [y \<leftarrow> xs. y = x]" | 
| 897 | by (induct xs) auto | |
| 15867 | 898 | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 899 | lemma nth_mem_multiset_of: "i < length ls \<Longrightarrow> (ls ! i) :# multiset_of ls" | 
| 26178 | 900 | apply (induct ls arbitrary: i) | 
| 901 | apply simp | |
| 902 | apply (case_tac i) | |
| 903 | apply auto | |
| 904 | done | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 905 | |
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 906 | lemma multiset_of_remove1: "multiset_of (remove1 a xs) = multiset_of xs - {#a#}"
 | 
| 26178 | 907 | by (induct xs) (auto simp add: multiset_eq_conv_count_eq) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 908 | |
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 909 | lemma multiset_of_eq_length: | 
| 26178 | 910 | assumes "multiset_of xs = multiset_of ys" | 
| 911 | shows "length xs = length ys" | |
| 912 | using assms | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 913 | proof (induct arbitrary: ys rule: length_induct) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 914 | case (1 xs ys) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 915 | show ?case | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 916 | proof (cases xs) | 
| 26145 | 917 | case Nil with "1.prems" show ?thesis by simp | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 918 | next | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 919 | case (Cons x xs') | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 920 | note xCons = Cons | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 921 | show ?thesis | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 922 | proof (cases ys) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 923 | case Nil | 
| 26145 | 924 | with "1.prems" Cons show ?thesis by simp | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 925 | next | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 926 | case (Cons y ys') | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 927 | have x_in_ys: "x = y \<or> x \<in> set ys'" | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 928 | proof (cases "x = y") | 
| 26145 | 929 | case True then show ?thesis .. | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 930 | next | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 931 | case False | 
| 26145 | 932 | 	from "1.prems" [symmetric] xCons Cons have "x :# multiset_of ys' + {#y#}" by simp
 | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 933 | with False show ?thesis by (simp add: mem_set_multiset_eq) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 934 | qed | 
| 26145 | 935 | from "1.hyps" have IH: "length xs' < length xs \<longrightarrow> | 
| 936 | (\<forall>x. multiset_of xs' = multiset_of x \<longrightarrow> length xs' = length x)" by blast | |
| 937 | from "1.prems" x_in_ys Cons xCons have "multiset_of xs' = multiset_of (remove1 x (y#ys'))" | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 938 | apply - | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 939 | apply (simp add: multiset_of_remove1, simp only: add_eq_conv_diff) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 940 | apply fastsimp | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 941 | done | 
| 26145 | 942 | with IH xCons have IH': "length xs' = length (remove1 x (y#ys'))" by fastsimp | 
| 943 | from x_in_ys have "x \<noteq> y \<Longrightarrow> length ys' > 0" by auto | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 944 | with Cons xCons x_in_ys IH' show ?thesis by (auto simp add: length_remove1) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 945 | qed | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 946 | qed | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 947 | qed | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 948 | |
| 26145 | 949 | text {*
 | 
| 950 | This lemma shows which properties suffice to show that a function | |
| 951 |   @{text "f"} with @{text "f xs = ys"} behaves like sort.
 | |
| 952 | *} | |
| 953 | lemma properties_for_sort: | |
| 954 | "multiset_of ys = multiset_of xs \<Longrightarrow> sorted ys \<Longrightarrow> sort xs = ys" | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 955 | proof (induct xs arbitrary: ys) | 
| 26145 | 956 | case Nil then show ?case by simp | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 957 | next | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 958 | case (Cons x xs) | 
| 26145 | 959 | then have "x \<in> set ys" | 
| 960 | by (auto simp add: mem_set_multiset_eq intro!: ccontr) | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 961 | with Cons.prems Cons.hyps [of "remove1 x ys"] show ?case | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 962 | by (simp add: sorted_remove1 multiset_of_remove1 insort_remove1) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 963 | qed | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 964 | |
| 15867 | 965 | |
| 15072 | 966 | subsection {* Pointwise ordering induced by count *}
 | 
| 967 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 968 | definition mset_le :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<le>#" 50) where | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 969 | [code del]: "A \<le># B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)" | 
| 26145 | 970 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 971 | definition mset_less :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<#" 50) where | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 972 | [code del]: "A <# B \<longleftrightarrow> A \<le># B \<and> A \<noteq> B" | 
| 25610 | 973 | |
| 26145 | 974 | notation mset_le (infix "\<subseteq>#" 50) | 
| 975 | notation mset_less (infix "\<subset>#" 50) | |
| 15072 | 976 | |
| 23611 | 977 | lemma mset_le_refl[simp]: "A \<le># A" | 
| 26178 | 978 | unfolding mset_le_def by auto | 
| 15072 | 979 | |
| 26145 | 980 | lemma mset_le_trans: "A \<le># B \<Longrightarrow> B \<le># C \<Longrightarrow> A \<le># C" | 
| 26178 | 981 | unfolding mset_le_def by (fast intro: order_trans) | 
| 15072 | 982 | |
| 26145 | 983 | lemma mset_le_antisym: "A \<le># B \<Longrightarrow> B \<le># A \<Longrightarrow> A = B" | 
| 26178 | 984 | apply (unfold mset_le_def) | 
| 985 | apply (rule multiset_eq_conv_count_eq [THEN iffD2]) | |
| 986 | apply (blast intro: order_antisym) | |
| 987 | done | |
| 15072 | 988 | |
| 26145 | 989 | lemma mset_le_exists_conv: "(A \<le># B) = (\<exists>C. B = A + C)" | 
| 26178 | 990 | apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI) | 
| 991 | apply (auto intro: multiset_eq_conv_count_eq [THEN iffD2]) | |
| 992 | done | |
| 15072 | 993 | |
| 23611 | 994 | lemma mset_le_mono_add_right_cancel[simp]: "(A + C \<le># B + C) = (A \<le># B)" | 
| 26178 | 995 | unfolding mset_le_def by auto | 
| 15072 | 996 | |
| 23611 | 997 | lemma mset_le_mono_add_left_cancel[simp]: "(C + A \<le># C + B) = (A \<le># B)" | 
| 26178 | 998 | unfolding mset_le_def by auto | 
| 15072 | 999 | |
| 23611 | 1000 | lemma mset_le_mono_add: "\<lbrakk> A \<le># B; C \<le># D \<rbrakk> \<Longrightarrow> A + C \<le># B + D" | 
| 26178 | 1001 | apply (unfold mset_le_def) | 
| 1002 | apply auto | |
| 1003 | apply (erule_tac x = a in allE)+ | |
| 1004 | apply auto | |
| 1005 | done | |
| 15072 | 1006 | |
| 23611 | 1007 | lemma mset_le_add_left[simp]: "A \<le># A + B" | 
| 26178 | 1008 | unfolding mset_le_def by auto | 
| 15072 | 1009 | |
| 23611 | 1010 | lemma mset_le_add_right[simp]: "B \<le># A + B" | 
| 26178 | 1011 | unfolding mset_le_def by auto | 
| 15072 | 1012 | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1013 | lemma mset_le_single: "a :# B \<Longrightarrow> {#a#} \<le># B"
 | 
| 26178 | 1014 | by (simp add: mset_le_def) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1015 | |
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1016 | lemma multiset_diff_union_assoc: "C \<le># B \<Longrightarrow> A + B - C = A + (B - C)" | 
| 26178 | 1017 | by (simp add: multiset_eq_conv_count_eq mset_le_def) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1018 | |
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1019 | lemma mset_le_multiset_union_diff_commute: | 
| 26178 | 1020 | assumes "B \<le># A" | 
| 1021 | shows "A - B + C = A + C - B" | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1022 | proof - | 
| 26145 | 1023 | from mset_le_exists_conv [of "B" "A"] assms have "\<exists>D. A = B + D" .. | 
| 1024 | from this obtain D where "A = B + D" .. | |
| 1025 | then show ?thesis | |
| 1026 | apply simp | |
| 1027 | apply (subst union_commute) | |
| 1028 | apply (subst multiset_diff_union_assoc) | |
| 1029 | apply simp | |
| 1030 | apply (simp add: diff_cancel) | |
| 1031 | apply (subst union_assoc) | |
| 1032 | apply (subst union_commute[of "B" _]) | |
| 1033 | apply (subst multiset_diff_union_assoc) | |
| 1034 | apply simp | |
| 1035 | apply (simp add: diff_cancel) | |
| 1036 | done | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1037 | qed | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1038 | |
| 23611 | 1039 | lemma multiset_of_remdups_le: "multiset_of (remdups xs) \<le># multiset_of xs" | 
| 26178 | 1040 | apply (induct xs) | 
| 1041 | apply auto | |
| 1042 | apply (rule mset_le_trans) | |
| 1043 | apply auto | |
| 1044 | done | |
| 23611 | 1045 | |
| 26145 | 1046 | lemma multiset_of_update: | 
| 1047 |   "i < length ls \<Longrightarrow> multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}"
 | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1048 | proof (induct ls arbitrary: i) | 
| 26145 | 1049 | case Nil then show ?case by simp | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1050 | next | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1051 | case (Cons x xs) | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1052 | show ?case | 
| 26145 | 1053 | proof (cases i) | 
| 1054 | case 0 then show ?thesis by simp | |
| 1055 | next | |
| 1056 | case (Suc i') | |
| 1057 | with Cons show ?thesis | |
| 1058 | apply simp | |
| 1059 | apply (subst union_assoc) | |
| 1060 |       apply (subst union_commute [where M = "{#v#}" and N = "{#x#}"])
 | |
| 1061 | apply (subst union_assoc [symmetric]) | |
| 1062 | apply simp | |
| 1063 | apply (rule mset_le_multiset_union_diff_commute) | |
| 1064 | apply (simp add: mset_le_single nth_mem_multiset_of) | |
| 1065 | done | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1066 | qed | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1067 | qed | 
| 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1068 | |
| 26145 | 1069 | lemma multiset_of_swap: | 
| 1070 | "i < length ls \<Longrightarrow> j < length ls \<Longrightarrow> | |
| 1071 | multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls" | |
| 26178 | 1072 | apply (case_tac "i = j") | 
| 1073 | apply simp | |
| 1074 | apply (simp add: multiset_of_update) | |
| 1075 | apply (subst elem_imp_eq_diff_union[symmetric]) | |
| 1076 | apply (simp add: nth_mem_multiset_of) | |
| 1077 | apply simp | |
| 1078 | done | |
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 1079 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
30663diff
changeset | 1080 | interpretation mset_order: order "op \<le>#" "op <#" | 
| 27682 | 1081 | proof qed (auto intro: order.intro mset_le_refl mset_le_antisym | 
| 1082 | mset_le_trans simp: mset_less_def) | |
| 23611 | 1083 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
30663diff
changeset | 1084 | interpretation mset_order_cancel_semigroup: | 
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29252diff
changeset | 1085 | pordered_cancel_ab_semigroup_add "op +" "op \<le>#" "op <#" | 
| 27682 | 1086 | proof qed (erule mset_le_mono_add [OF mset_le_refl]) | 
| 23611 | 1087 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
30663diff
changeset | 1088 | interpretation mset_order_semigroup_cancel: | 
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29252diff
changeset | 1089 | pordered_ab_semigroup_add_imp_le "op +" "op \<le>#" "op <#" | 
| 27682 | 1090 | proof qed simp | 
| 15072 | 1091 | |
| 25610 | 1092 | |
| 26145 | 1093 | lemma mset_lessD: "A \<subset># B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B" | 
| 26178 | 1094 | apply (clarsimp simp: mset_le_def mset_less_def) | 
| 1095 | apply (erule_tac x=x in allE) | |
| 1096 | apply auto | |
| 1097 | done | |
| 25610 | 1098 | |
| 26145 | 1099 | lemma mset_leD: "A \<subseteq># B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B" | 
| 26178 | 1100 | apply (clarsimp simp: mset_le_def mset_less_def) | 
| 1101 | apply (erule_tac x = x in allE) | |
| 1102 | apply auto | |
| 1103 | done | |
| 25610 | 1104 | |
| 26145 | 1105 | lemma mset_less_insertD: "(A + {#x#} \<subset># B) \<Longrightarrow> (x \<in># B \<and> A \<subset># B)"
 | 
| 26178 | 1106 | apply (rule conjI) | 
| 1107 | apply (simp add: mset_lessD) | |
| 1108 | apply (clarsimp simp: mset_le_def mset_less_def) | |
| 1109 | apply safe | |
| 1110 | apply (erule_tac x = a in allE) | |
| 1111 | apply (auto split: split_if_asm) | |
| 1112 | done | |
| 25610 | 1113 | |
| 26145 | 1114 | lemma mset_le_insertD: "(A + {#x#} \<subseteq># B) \<Longrightarrow> (x \<in># B \<and> A \<subseteq># B)"
 | 
| 26178 | 1115 | apply (rule conjI) | 
| 1116 | apply (simp add: mset_leD) | |
| 1117 | apply (force simp: mset_le_def mset_less_def split: split_if_asm) | |
| 1118 | done | |
| 25610 | 1119 | |
| 1120 | lemma mset_less_of_empty[simp]: "A \<subset># {#} = False" 
 | |
| 26178 | 1121 | by (induct A) (auto simp: mset_le_def mset_less_def) | 
| 25610 | 1122 | |
| 1123 | lemma multi_psub_of_add_self[simp]: "A \<subset># A + {#x#}"
 | |
| 26178 | 1124 | by (auto simp: mset_le_def mset_less_def) | 
| 25610 | 1125 | |
| 1126 | lemma multi_psub_self[simp]: "A \<subset># A = False" | |
| 26178 | 1127 | by (auto simp: mset_le_def mset_less_def) | 
| 25610 | 1128 | |
| 1129 | lemma mset_less_add_bothsides: | |
| 1130 |   "T + {#x#} \<subset># S + {#x#} \<Longrightarrow> T \<subset># S"
 | |
| 26178 | 1131 | by (auto simp: mset_le_def mset_less_def) | 
| 25610 | 1132 | |
| 1133 | lemma mset_less_empty_nonempty: "({#} \<subset># S) = (S \<noteq> {#})"
 | |
| 26178 | 1134 | by (auto simp: mset_le_def mset_less_def) | 
| 25610 | 1135 | |
| 1136 | lemma mset_less_size: "A \<subset># B \<Longrightarrow> size A < size B" | |
| 1137 | proof (induct A arbitrary: B) | |
| 1138 | case (empty M) | |
| 26145 | 1139 |   then have "M \<noteq> {#}" by (simp add: mset_less_empty_nonempty)
 | 
| 25610 | 1140 |   then obtain M' x where "M = M' + {#x#}" 
 | 
| 1141 | by (blast dest: multi_nonempty_split) | |
| 26145 | 1142 | then show ?case by simp | 
| 25610 | 1143 | next | 
| 1144 | case (add S x T) | |
| 1145 | have IH: "\<And>B. S \<subset># B \<Longrightarrow> size S < size B" by fact | |
| 1146 |   have SxsubT: "S + {#x#} \<subset># T" by fact
 | |
| 26145 | 1147 | then have "x \<in># T" and "S \<subset># T" by (auto dest: mset_less_insertD) | 
| 25610 | 1148 |   then obtain T' where T: "T = T' + {#x#}" 
 | 
| 1149 | by (blast dest: multi_member_split) | |
| 26145 | 1150 | then have "S \<subset># T'" using SxsubT | 
| 25610 | 1151 | by (blast intro: mset_less_add_bothsides) | 
| 26145 | 1152 | then have "size S < size T'" using IH by simp | 
| 1153 | then show ?case using T by simp | |
| 25610 | 1154 | qed | 
| 1155 | ||
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29252diff
changeset | 1156 | lemmas mset_less_trans = mset_order.less_trans | 
| 25610 | 1157 | |
| 1158 | lemma mset_less_diff_self: "c \<in># B \<Longrightarrow> B - {#c#} \<subset># B"
 | |
| 26178 | 1159 | by (auto simp: mset_le_def mset_less_def multi_drop_mem_not_eq) | 
| 25610 | 1160 | |
| 26145 | 1161 | |
| 25610 | 1162 | subsection {* Strong induction and subset induction for multisets *}
 | 
| 1163 | ||
| 26016 | 1164 | text {* Well-foundedness of proper subset operator: *}
 | 
| 25610 | 1165 | |
| 26145 | 1166 | text {* proper multiset subset *}
 | 
| 25610 | 1167 | definition | 
| 26145 | 1168 |   mset_less_rel :: "('a multiset * 'a multiset) set" where
 | 
| 1169 |   "mset_less_rel = {(A,B). A \<subset># B}"
 | |
| 25610 | 1170 | |
| 1171 | lemma multiset_add_sub_el_shuffle: | |
| 26145 | 1172 | assumes "c \<in># B" and "b \<noteq> c" | 
| 25610 | 1173 |   shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}"
 | 
| 1174 | proof - | |
| 26145 | 1175 |   from `c \<in># B` obtain A where B: "B = A + {#c#}" 
 | 
| 25610 | 1176 | by (blast dest: multi_member_split) | 
| 1177 |   have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp
 | |
| 26145 | 1178 |   then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" 
 | 
| 25610 | 1179 | by (simp add: union_ac) | 
| 26145 | 1180 | then show ?thesis using B by simp | 
| 25610 | 1181 | qed | 
| 1182 | ||
| 1183 | lemma wf_mset_less_rel: "wf mset_less_rel" | |
| 26178 | 1184 | apply (unfold mset_less_rel_def) | 
| 1185 | apply (rule wf_measure [THEN wf_subset, where f1=size]) | |
| 1186 | apply (clarsimp simp: measure_def inv_image_def mset_less_size) | |
| 1187 | done | |
| 25610 | 1188 | |
| 26016 | 1189 | text {* The induction rules: *}
 | 
| 25610 | 1190 | |
| 1191 | lemma full_multiset_induct [case_names less]: | |
| 26178 | 1192 | assumes ih: "\<And>B. \<forall>A. A \<subset># B \<longrightarrow> P A \<Longrightarrow> P B" | 
| 1193 | shows "P B" | |
| 1194 | apply (rule wf_mset_less_rel [THEN wf_induct]) | |
| 1195 | apply (rule ih, auto simp: mset_less_rel_def) | |
| 1196 | done | |
| 25610 | 1197 | |
| 1198 | lemma multi_subset_induct [consumes 2, case_names empty add]: | |
| 26178 | 1199 | assumes "F \<subseteq># A" | 
| 1200 |   and empty: "P {#}"
 | |
| 1201 |   and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (F + {#a#})"
 | |
| 1202 | shows "P F" | |
| 25610 | 1203 | proof - | 
| 1204 | from `F \<subseteq># A` | |
| 1205 | show ?thesis | |
| 1206 | proof (induct F) | |
| 1207 |     show "P {#}" by fact
 | |
| 1208 | next | |
| 1209 | fix x F | |
| 1210 |     assume P: "F \<subseteq># A \<Longrightarrow> P F" and i: "F + {#x#} \<subseteq># A"
 | |
| 1211 |     show "P (F + {#x#})"
 | |
| 1212 | proof (rule insert) | |
| 1213 | from i show "x \<in># A" by (auto dest: mset_le_insertD) | |
| 26145 | 1214 | from i have "F \<subseteq># A" by (auto dest: mset_le_insertD) | 
| 25610 | 1215 | with P show "P F" . | 
| 1216 | qed | |
| 1217 | qed | |
| 1218 | qed | |
| 1219 | ||
| 26016 | 1220 | text{* A consequence: Extensionality. *}
 | 
| 25610 | 1221 | |
| 26145 | 1222 | lemma multi_count_eq: "(\<forall>x. count A x = count B x) = (A = B)" | 
| 26178 | 1223 | apply (rule iffI) | 
| 1224 | prefer 2 | |
| 1225 | apply clarsimp | |
| 1226 | apply (induct A arbitrary: B rule: full_multiset_induct) | |
| 1227 | apply (rename_tac C) | |
| 1228 | apply (case_tac B rule: multiset_cases) | |
| 1229 | apply (simp add: empty_multiset_count) | |
| 1230 | apply simp | |
| 1231 | apply (case_tac "x \<in># C") | |
| 1232 | apply (force dest: multi_member_split) | |
| 1233 | apply (erule_tac x = x in allE) | |
| 1234 | apply simp | |
| 1235 | done | |
| 25610 | 1236 | |
| 1237 | lemmas multi_count_ext = multi_count_eq [THEN iffD1, rule_format] | |
| 1238 | ||
| 26145 | 1239 | |
| 25610 | 1240 | subsection {* The fold combinator *}
 | 
| 1241 | ||
| 26145 | 1242 | text {*
 | 
| 1243 | The intended behaviour is | |
| 1244 |   @{text "fold_mset f z {#x\<^isub>1, ..., x\<^isub>n#} = f x\<^isub>1 (\<dots> (f x\<^isub>n z)\<dots>)"}
 | |
| 1245 |   if @{text f} is associative-commutative. 
 | |
| 25610 | 1246 | *} | 
| 1247 | ||
| 26145 | 1248 | text {*
 | 
| 1249 |   The graph of @{text "fold_mset"}, @{text "z"}: the start element,
 | |
| 1250 |   @{text "f"}: folding function, @{text "A"}: the multiset, @{text
 | |
| 1251 | "y"}: the result. | |
| 1252 | *} | |
| 25610 | 1253 | inductive | 
| 25759 | 1254 |   fold_msetG :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b \<Rightarrow> bool" 
 | 
| 25610 | 1255 | for f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" | 
| 1256 | and z :: 'b | |
| 1257 | where | |
| 25759 | 1258 |   emptyI [intro]:  "fold_msetG f z {#} z"
 | 
| 1259 | | insertI [intro]: "fold_msetG f z A y \<Longrightarrow> fold_msetG f z (A + {#x#}) (f x y)"
 | |
| 25610 | 1260 | |
| 25759 | 1261 | inductive_cases empty_fold_msetGE [elim!]: "fold_msetG f z {#} x"
 | 
| 1262 | inductive_cases insert_fold_msetGE: "fold_msetG f z (A + {#}) y" 
 | |
| 25610 | 1263 | |
| 1264 | definition | |
| 26145 | 1265 |   fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b" where
 | 
| 1266 | "fold_mset f z A = (THE x. fold_msetG f z A x)" | |
| 25610 | 1267 | |
| 25759 | 1268 | lemma Diff1_fold_msetG: | 
| 26145 | 1269 |   "fold_msetG f z (A - {#x#}) y \<Longrightarrow> x \<in># A \<Longrightarrow> fold_msetG f z A (f x y)"
 | 
| 26178 | 1270 | apply (frule_tac x = x in fold_msetG.insertI) | 
| 1271 | apply auto | |
| 1272 | done | |
| 25610 | 1273 | |
| 25759 | 1274 | lemma fold_msetG_nonempty: "\<exists>x. fold_msetG f z A x" | 
| 26178 | 1275 | apply (induct A) | 
| 1276 | apply blast | |
| 1277 | apply clarsimp | |
| 1278 | apply (drule_tac x = x in fold_msetG.insertI) | |
| 1279 | apply auto | |
| 1280 | done | |
| 25610 | 1281 | |
| 25759 | 1282 | lemma fold_mset_empty[simp]: "fold_mset f z {#} = z"
 | 
| 26178 | 1283 | unfolding fold_mset_def by blast | 
| 25610 | 1284 | |
| 1285 | locale left_commutative = | |
| 26178 | 1286 | fixes f :: "'a => 'b => 'b" | 
| 1287 | assumes left_commute: "f x (f y z) = f y (f x z)" | |
| 26145 | 1288 | begin | 
| 25610 | 1289 | |
| 26145 | 1290 | lemma fold_msetG_determ: | 
| 1291 | "fold_msetG f z A x \<Longrightarrow> fold_msetG f z A y \<Longrightarrow> y = x" | |
| 25610 | 1292 | proof (induct arbitrary: x y z rule: full_multiset_induct) | 
| 1293 | case (less M x\<^isub>1 x\<^isub>2 Z) | |
| 1294 | have IH: "\<forall>A. A \<subset># M \<longrightarrow> | |
| 25759 | 1295 | (\<forall>x x' x''. fold_msetG f x'' A x \<longrightarrow> fold_msetG f x'' A x' | 
| 25610 | 1296 | \<longrightarrow> x' = x)" by fact | 
| 25759 | 1297 | have Mfoldx\<^isub>1: "fold_msetG f Z M x\<^isub>1" and Mfoldx\<^isub>2: "fold_msetG f Z M x\<^isub>2" by fact+ | 
| 25610 | 1298 | show ?case | 
| 25759 | 1299 | proof (rule fold_msetG.cases [OF Mfoldx\<^isub>1]) | 
| 25610 | 1300 |     assume "M = {#}" and "x\<^isub>1 = Z"
 | 
| 26145 | 1301 | then show ?case using Mfoldx\<^isub>2 by auto | 
| 25610 | 1302 | next | 
| 1303 | fix B b u | |
| 25759 | 1304 |     assume "M = B + {#b#}" and "x\<^isub>1 = f b u" and Bu: "fold_msetG f Z B u"
 | 
| 26145 | 1305 |     then have MBb: "M = B + {#b#}" and x\<^isub>1: "x\<^isub>1 = f b u" by auto
 | 
| 25610 | 1306 | show ?case | 
| 25759 | 1307 | proof (rule fold_msetG.cases [OF Mfoldx\<^isub>2]) | 
| 25610 | 1308 |       assume "M = {#}" "x\<^isub>2 = Z"
 | 
| 26145 | 1309 | then show ?case using Mfoldx\<^isub>1 by auto | 
| 25610 | 1310 | next | 
| 1311 | fix C c v | |
| 25759 | 1312 |       assume "M = C + {#c#}" and "x\<^isub>2 = f c v" and Cv: "fold_msetG f Z C v"
 | 
| 26145 | 1313 |       then have MCc: "M = C + {#c#}" and x\<^isub>2: "x\<^isub>2 = f c v" by auto
 | 
| 1314 | then have CsubM: "C \<subset># M" by simp | |
| 25610 | 1315 | from MBb have BsubM: "B \<subset># M" by simp | 
| 1316 | show ?case | |
| 1317 | proof cases | |
| 1318 | assume "b=c" | |
| 1319 | then moreover have "B = C" using MBb MCc by auto | |
| 1320 | ultimately show ?thesis using Bu Cv x\<^isub>1 x\<^isub>2 CsubM IH by auto | |
| 1321 | next | |
| 1322 | assume diff: "b \<noteq> c" | |
| 1323 |         let ?D = "B - {#c#}"
 | |
| 1324 | have cinB: "c \<in># B" and binC: "b \<in># C" using MBb MCc diff | |
| 1325 | by (auto intro: insert_noteq_member dest: sym) | |
| 1326 |         have "B - {#c#} \<subset># B" using cinB by (rule mset_less_diff_self)
 | |
| 26145 | 1327 | then have DsubM: "?D \<subset># M" using BsubM by (blast intro: mset_less_trans) | 
| 25610 | 1328 |         from MBb MCc have "B + {#b#} = C + {#c#}" by blast
 | 
| 26145 | 1329 |         then have [simp]: "B + {#b#} - {#c#} = C"
 | 
| 25610 | 1330 | using MBb MCc binC cinB by auto | 
| 1331 |         have B: "B = ?D + {#c#}" and C: "C = ?D + {#b#}"
 | |
| 1332 | using MBb MCc diff binC cinB | |
| 1333 | by (auto simp: multiset_add_sub_el_shuffle) | |
| 25759 | 1334 | then obtain d where Dfoldd: "fold_msetG f Z ?D d" | 
| 1335 | using fold_msetG_nonempty by iprover | |
| 26145 | 1336 | then have "fold_msetG f Z B (f c d)" using cinB | 
| 25759 | 1337 | by (rule Diff1_fold_msetG) | 
| 26145 | 1338 | then have "f c d = u" using IH BsubM Bu by blast | 
| 25610 | 1339 | moreover | 
| 25759 | 1340 | have "fold_msetG f Z C (f b d)" using binC cinB diff Dfoldd | 
| 25610 | 1341 | by (auto simp: multiset_add_sub_el_shuffle | 
| 25759 | 1342 | dest: fold_msetG.insertI [where x=b]) | 
| 26145 | 1343 | then have "f b d = v" using IH CsubM Cv by blast | 
| 25610 | 1344 | ultimately show ?thesis using x\<^isub>1 x\<^isub>2 | 
| 1345 | by (auto simp: left_commute) | |
| 1346 | qed | |
| 1347 | qed | |
| 1348 | qed | |
| 1349 | qed | |
| 1350 | ||
| 26145 | 1351 | lemma fold_mset_insert_aux: | 
| 1352 |   "(fold_msetG f z (A + {#x#}) v) =
 | |
| 25759 | 1353 | (\<exists>y. fold_msetG f z A y \<and> v = f x y)" | 
| 26178 | 1354 | apply (rule iffI) | 
| 1355 | prefer 2 | |
| 1356 | apply blast | |
| 1357 | apply (rule_tac A=A and f=f in fold_msetG_nonempty [THEN exE, standard]) | |
| 1358 | apply (blast intro: fold_msetG_determ) | |
| 1359 | done | |
| 25610 | 1360 | |
| 26145 | 1361 | lemma fold_mset_equality: "fold_msetG f z A y \<Longrightarrow> fold_mset f z A = y" | 
| 26178 | 1362 | unfolding fold_mset_def by (blast intro: fold_msetG_determ) | 
| 25610 | 1363 | |
| 26145 | 1364 | lemma fold_mset_insert: | 
| 26178 | 1365 |   "fold_mset f z (A + {#x#}) = f x (fold_mset f z A)"
 | 
| 1366 | apply (simp add: fold_mset_def fold_mset_insert_aux union_commute) | |
| 1367 | apply (rule the_equality) | |
| 1368 | apply (auto cong add: conj_cong | |
| 26145 | 1369 | simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) | 
| 26178 | 1370 | done | 
| 25759 | 1371 | |
| 26145 | 1372 | lemma fold_mset_insert_idem: | 
| 26178 | 1373 |   "fold_mset f z (A + {#a#}) = f a (fold_mset f z A)"
 | 
| 1374 | apply (simp add: fold_mset_def fold_mset_insert_aux) | |
| 1375 | apply (rule the_equality) | |
| 1376 | apply (auto cong add: conj_cong | |
| 26145 | 1377 | simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) | 
| 26178 | 1378 | done | 
| 25610 | 1379 | |
| 26145 | 1380 | lemma fold_mset_commute: "f x (fold_mset f z A) = fold_mset f (f x z) A" | 
| 26178 | 1381 | by (induct A) (auto simp: fold_mset_insert left_commute [of x]) | 
| 1382 | ||
| 26145 | 1383 | lemma fold_mset_single [simp]: "fold_mset f z {#x#} = f x z"
 | 
| 26178 | 1384 | using fold_mset_insert [of z "{#}"] by simp
 | 
| 25610 | 1385 | |
| 26145 | 1386 | lemma fold_mset_union [simp]: | 
| 1387 | "fold_mset f z (A+B) = fold_mset f (fold_mset f z A) B" | |
| 25759 | 1388 | proof (induct A) | 
| 26145 | 1389 | case empty then show ?case by simp | 
| 25759 | 1390 | next | 
| 26145 | 1391 | case (add A x) | 
| 1392 |   have "A + {#x#} + B = (A+B) + {#x#}" by(simp add:union_ac)
 | |
| 1393 |   then have "fold_mset f z (A + {#x#} + B) = f x (fold_mset f z (A + B))" 
 | |
| 1394 | by (simp add: fold_mset_insert) | |
| 1395 |   also have "\<dots> = fold_mset f (fold_mset f z (A + {#x#})) B"
 | |
| 1396 | by (simp add: fold_mset_commute[of x,symmetric] add fold_mset_insert) | |
| 1397 | finally show ?case . | |
| 25759 | 1398 | qed | 
| 1399 | ||
| 26145 | 1400 | lemma fold_mset_fusion: | 
| 27611 | 1401 | assumes "left_commutative g" | 
| 1402 | shows "(\<And>x y. h (g x y) = f x (h y)) \<Longrightarrow> h (fold_mset g w A) = fold_mset f (h w) A" (is "PROP ?P") | |
| 1403 | proof - | |
| 29233 | 1404 | interpret left_commutative g by fact | 
| 27611 | 1405 | show "PROP ?P" by (induct A) auto | 
| 1406 | qed | |
| 25610 | 1407 | |
| 26145 | 1408 | lemma fold_mset_rec: | 
| 1409 | assumes "a \<in># A" | |
| 25759 | 1410 |   shows "fold_mset f z A = f a (fold_mset f z (A - {#a#}))"
 | 
| 25610 | 1411 | proof - | 
| 26145 | 1412 |   from assms obtain A' where "A = A' + {#a#}"
 | 
| 1413 | by (blast dest: multi_member_split) | |
| 1414 | then show ?thesis by simp | |
| 25610 | 1415 | qed | 
| 1416 | ||
| 26145 | 1417 | end | 
| 1418 | ||
| 1419 | text {*
 | |
| 1420 | A note on code generation: When defining some function containing a | |
| 1421 |   subterm @{term"fold_mset F"}, code generation is not automatic. When
 | |
| 1422 |   interpreting locale @{text left_commutative} with @{text F}, the
 | |
| 1423 |   would be code thms for @{const fold_mset} become thms like
 | |
| 1424 |   @{term"fold_mset F z {#} = z"} where @{text F} is not a pattern but
 | |
| 1425 | contains defined symbols, i.e.\ is not a code thm. Hence a separate | |
| 1426 |   constant with its own code thms needs to be introduced for @{text
 | |
| 1427 | F}. See the image operator below. | |
| 1428 | *} | |
| 1429 | ||
| 26016 | 1430 | |
| 1431 | subsection {* Image *}
 | |
| 1432 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1433 | definition [code del]: | 
| 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1434 |  "image_mset f = fold_mset (op + o single o f) {#}"
 | 
| 26016 | 1435 | |
| 30729 
461ee3e49ad3
interpretation/interpret: prefixes are mandatory by default;
 wenzelm parents: 
30663diff
changeset | 1436 | interpretation image_left_comm: left_commutative "op + o single o f" | 
| 28823 | 1437 | proof qed (simp add:union_ac) | 
| 26016 | 1438 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1439 | lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
 | 
| 26178 | 1440 | by (simp add: image_mset_def) | 
| 26016 | 1441 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1442 | lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}"
 | 
| 26178 | 1443 | by (simp add: image_mset_def) | 
| 26016 | 1444 | |
| 1445 | lemma image_mset_insert: | |
| 1446 |   "image_mset f (M + {#a#}) = image_mset f M + {#f a#}"
 | |
| 26178 | 1447 | by (simp add: image_mset_def add_ac) | 
| 26016 | 1448 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1449 | lemma image_mset_union [simp]: | 
| 26016 | 1450 | "image_mset f (M+N) = image_mset f M + image_mset f N" | 
| 26178 | 1451 | apply (induct N) | 
| 1452 | apply simp | |
| 1453 | apply (simp add: union_assoc [symmetric] image_mset_insert) | |
| 1454 | done | |
| 26016 | 1455 | |
| 26145 | 1456 | lemma size_image_mset [simp]: "size (image_mset f M) = size M" | 
| 26178 | 1457 | by (induct M) simp_all | 
| 26016 | 1458 | |
| 26145 | 1459 | lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \<longleftrightarrow> M = {#}"
 | 
| 26178 | 1460 | by (cases M) auto | 
| 26016 | 1461 | |
| 26145 | 1462 | syntax | 
| 1463 | comprehension1_mset :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" | |
| 1464 |       ("({#_/. _ :# _#})")
 | |
| 1465 | translations | |
| 1466 |   "{#e. x:#M#}" == "CONST image_mset (%x. e) M"
 | |
| 26016 | 1467 | |
| 26145 | 1468 | syntax | 
| 1469 | comprehension2_mset :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" | |
| 1470 |       ("({#_/ | _ :# _./ _#})")
 | |
| 26016 | 1471 | translations | 
| 26033 | 1472 |   "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}"
 | 
| 26016 | 1473 | |
| 26145 | 1474 | text {*
 | 
| 1475 |   This allows to write not just filters like @{term "{#x:#M. x<c#}"}
 | |
| 1476 |   but also images like @{term "{#x+x. x:#M #}"} and @{term [source]
 | |
| 1477 |   "{#x+x|x:#M. x<c#}"}, where the latter is currently displayed as
 | |
| 1478 |   @{term "{#x+x|x:#M. x<c#}"}.
 | |
| 1479 | *} | |
| 26016 | 1480 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1481 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1482 | subsection {* Termination proofs with multiset orders *}
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1483 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1484 | 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: 
28823diff
changeset | 1485 |   and multi_member_this: "x \<in># {# x #} + XS"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1486 |   and multi_member_last: "x \<in># {# x #}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1487 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1488 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1489 | definition "ms_strict = mult pair_less" | 
| 30428 | 1490 | definition [code del]: "ms_weak = ms_strict \<union> Id" | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1491 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1492 | 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: 
28823diff
changeset | 1493 | 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: 
28823diff
changeset | 1494 | 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: 
28823diff
changeset | 1495 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1496 | lemma smsI: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1497 | "(set_of A, set_of B) \<in> max_strict \<Longrightarrow> (Z + A, Z + B) \<in> ms_strict" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1498 | unfolding ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1499 | 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: 
28823diff
changeset | 1500 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1501 | lemma wmsI: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1502 |   "(set_of A, set_of B) \<in> max_strict \<or> A = {#} \<and> B = {#}
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1503 | \<Longrightarrow> (Z + A, Z + B) \<in> ms_weak" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1504 | unfolding ms_weak_def ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1505 | 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: 
28823diff
changeset | 1506 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1507 | inductive pw_leq | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1508 | where | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1509 |   pw_leq_empty: "pw_leq {#} {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1510 | | 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: 
28823diff
changeset | 1511 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1512 | lemma pw_leq_lstep: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1513 |   "(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: 
28823diff
changeset | 1514 | 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: 
28823diff
changeset | 1515 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1516 | lemma pw_leq_split: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1517 | assumes "pw_leq X Y" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1518 |   shows "\<exists>A B Z. X = A + Z \<and> Y = B + Z \<and> ((set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1519 | using assms | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1520 | proof (induct) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1521 | case pw_leq_empty thus ?case by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1522 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1523 | case (pw_leq_step x y X Y) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1524 | then obtain A B Z where | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1525 | [simp]: "X = A + Z" "Y = B + Z" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1526 |       and 1[simp]: "(set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#})" 
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1527 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1528 | from pw_leq_step have "x = y \<or> (x, y) \<in> pair_less" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1529 | unfolding pair_leq_def by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1530 | thus ?case | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1531 | proof | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1532 | assume [simp]: "x = y" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1533 | have | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1534 |       "{#x#} + X = A + ({#y#}+Z) 
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1535 |       \<and> {#y#} + Y = B + ({#y#}+Z)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1536 |       \<and> ((set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1537 | by (auto simp: add_ac) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1538 | thus ?case by (intro exI) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1539 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1540 | assume A: "(x, y) \<in> pair_less" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1541 |     let ?A' = "{#x#} + A" and ?B' = "{#y#} + B"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1542 |     have "{#x#} + X = ?A' + Z"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1543 |       "{#y#} + Y = ?B' + Z"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1544 | by (auto simp add: add_ac) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1545 | moreover have | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1546 | "(set_of ?A', set_of ?B') \<in> max_strict" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1547 | using 1 A unfolding max_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1548 | by (auto elim!: max_ext.cases) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1549 | ultimately show ?thesis by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1550 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1551 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1552 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1553 | lemma | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1554 | assumes pwleq: "pw_leq Z Z'" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1555 | shows ms_strictI: "(set_of A, set_of B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_strict" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1556 | and ms_weakI1: "(set_of A, set_of B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_weak" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1557 |   and   ms_weakI2:  "(Z + {#}, Z' + {#}) \<in> ms_weak"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1558 | proof - | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1559 | from pw_leq_split[OF pwleq] | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1560 | obtain A' B' Z'' | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1561 | where [simp]: "Z = A' + Z''" "Z' = B' + Z''" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1562 |     and mx_or_empty: "(set_of A', set_of B') \<in> max_strict \<or> (A' = {#} \<and> B' = {#})"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1563 | by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1564 |   {
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1565 | assume max: "(set_of A, set_of B) \<in> max_strict" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1566 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1567 | 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: 
28823diff
changeset | 1568 | proof | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1569 | assume max': "(set_of A', set_of B') \<in> max_strict" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1570 | with max have "(set_of (A + A'), set_of (B + B')) \<in> max_strict" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1571 | by (auto simp: max_strict_def intro: max_ext_additive) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1572 | thus ?thesis by (rule smsI) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1573 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1574 |       assume [simp]: "A' = {#} \<and> B' = {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1575 | show ?thesis by (rule smsI) (auto intro: max) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1576 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1577 | thus "(Z + A, Z' + B) \<in> ms_strict" by (simp add:add_ac) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1578 | 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: 
28823diff
changeset | 1579 | } | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1580 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1581 | have "(Z'' + A', Z'' + B') \<in> ms_weak" by (rule wmsI) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1582 |   thus "(Z + {#}, Z' + {#}) \<in> ms_weak" by (simp add:add_ac)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1583 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1584 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1585 | lemma empty_idemp: "{#} + x = x" "x + {#} = x"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1586 | and nonempty_plus: "{# x #} + rs \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1587 | and nonempty_single: "{# x #} \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1588 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1589 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1590 | setup {*
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1591 | let | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1592 |   fun msetT T = Type ("Multiset.multiset", [T]);
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1593 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1594 |   fun mk_mset T [] = Const (@{const_name Mempty}, msetT T)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1595 |     | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1596 | | mk_mset T (x :: xs) = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1597 |           Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1598 | mk_mset T [x] $ mk_mset T xs | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1599 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1600 | fun mset_member_tac m i = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1601 | (if m <= 0 then | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1602 |            rtac @{thm multi_member_this} i ORELSE rtac @{thm multi_member_last} i
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1603 | else | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1604 |            rtac @{thm multi_member_skip} i THEN mset_member_tac (m - 1) i)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1605 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1606 | val mset_nonempty_tac = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1607 |       rtac @{thm nonempty_plus} ORELSE' rtac @{thm nonempty_single}
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1608 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1609 | val regroup_munion_conv = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1610 |       FundefLib.regroup_conv @{const_name Multiset.Mempty} @{const_name plus}
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1611 |         (map (fn t => t RS eq_reflection) (@{thms union_ac} @ @{thms empty_idemp}))
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1612 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1613 | fun unfold_pwleq_tac i = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1614 |     (rtac @{thm pw_leq_step} i THEN (fn st => unfold_pwleq_tac (i + 1) st))
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1615 |       ORELSE (rtac @{thm pw_leq_lstep} i)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1616 |       ORELSE (rtac @{thm pw_leq_empty} i)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1617 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1618 |   val set_of_simps = [@{thm set_of_empty}, @{thm set_of_single}, @{thm set_of_union},
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1619 |                       @{thm Un_insert_left}, @{thm Un_empty_left}]
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1620 | in | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1621 | ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1622 |   {
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1623 | msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1624 | mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1625 | mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_of_simps, | 
| 30595 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30428diff
changeset | 1626 |     smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1},
 | 
| 
c87a3350f5a9
proper spacing before ML antiquotations -- note that @ may be part of symbolic ML identifiers;
 wenzelm parents: 
30428diff
changeset | 1627 |     reduction_pair= @{thm ms_reduction_pair}
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1628 | }) | 
| 10249 | 1629 | end | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1630 | *} | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1631 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1632 | end |