| author | wenzelm | 
| Sun, 12 Jun 2011 16:19:29 +0200 | |
| changeset 43368 | 0dc67b3cf8a5 | 
| parent 42871 | 1c0b99f950d9 | 
| child 44339 | eda6aef75939 | 
| 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 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 5 | header {* (Finite) multisets *}
 | 
| 10249 | 6 | |
| 15131 | 7 | theory Multiset | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 8 | imports Main | 
| 15131 | 9 | begin | 
| 10249 | 10 | |
| 11 | subsection {* The type of multisets *}
 | |
| 12 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 13 | typedef 'a multiset = "{f :: 'a => nat. finite {x. f x > 0}}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 14 | morphisms count Abs_multiset | 
| 10249 | 15 | proof | 
| 11464 | 16 | show "(\<lambda>x. 0::nat) \<in> ?multiset" by simp | 
| 10249 | 17 | qed | 
| 18 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 19 | lemmas multiset_typedef = Abs_multiset_inverse count_inverse count | 
| 19086 | 20 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 21 | abbreviation Melem :: "'a => 'a multiset => bool"  ("(_/ :# _)" [50, 51] 50) where
 | 
| 25610 | 22 | "a :# M == 0 < count M a" | 
| 23 | ||
| 26145 | 24 | notation (xsymbols) | 
| 25 | Melem (infix "\<in>#" 50) | |
| 10249 | 26 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 27 | lemma multiset_eq_iff: | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 28 | "M = N \<longleftrightarrow> (\<forall>a. count M a = count N a)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 29 | by (simp only: count_inject [symmetric] fun_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 30 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 31 | lemma multiset_eqI: | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 32 | "(\<And>x. count A x = count B x) \<Longrightarrow> A = B" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 33 | using multiset_eq_iff by auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 34 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 35 | text {*
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 36 |  \medskip Preservation of the representing set @{term multiset}.
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 37 | *} | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 38 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 39 | lemma const0_in_multiset: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 40 | "(\<lambda>a. 0) \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 41 | by (simp add: multiset_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 42 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 43 | lemma only1_in_multiset: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 44 | "(\<lambda>b. if b = a then n else 0) \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 45 | by (simp add: multiset_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 46 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 47 | lemma union_preserves_multiset: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 48 | "M \<in> multiset \<Longrightarrow> N \<in> multiset \<Longrightarrow> (\<lambda>a. M a + N a) \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 49 | by (simp add: multiset_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 50 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 51 | lemma diff_preserves_multiset: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 52 | assumes "M \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 53 | shows "(\<lambda>a. M a - N a) \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 54 | proof - | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 55 |   have "{x. N x < M x} \<subseteq> {x. 0 < M x}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 56 | by auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 57 | with assms show ?thesis | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 58 | by (auto simp add: multiset_def intro: finite_subset) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 59 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 60 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 61 | lemma filter_preserves_multiset: | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 62 | assumes "M \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 63 | shows "(\<lambda>x. if P x then M x else 0) \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 64 | proof - | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 65 |   have "{x. (P x \<longrightarrow> 0 < M x) \<and> P x} \<subseteq> {x. 0 < M x}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 66 | by auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 67 | with assms show ?thesis | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 68 | by (auto simp add: multiset_def intro: finite_subset) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 69 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 70 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 71 | lemmas in_multiset = const0_in_multiset only1_in_multiset | 
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 72 | union_preserves_multiset diff_preserves_multiset filter_preserves_multiset | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 73 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 74 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 75 | subsection {* Representing multisets *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 76 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 77 | text {* Multiset enumeration *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 78 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 79 | instantiation multiset :: (type) "{zero, plus}"
 | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 80 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 81 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 82 | definition Mempty_def: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 83 | "0 = Abs_multiset (\<lambda>a. 0)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 84 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 85 | abbreviation Mempty :: "'a multiset" ("{#}") where
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 86 | "Mempty \<equiv> 0" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 87 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 88 | definition union_def: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 89 | "M + N = Abs_multiset (\<lambda>a. count M a + count N a)" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 90 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 91 | instance .. | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 92 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 93 | end | 
| 10249 | 94 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 95 | definition single :: "'a => 'a multiset" where | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 96 | "single a = Abs_multiset (\<lambda>b. if b = a then 1 else 0)" | 
| 15869 | 97 | |
| 26145 | 98 | syntax | 
| 26176 | 99 |   "_multiset" :: "args => 'a multiset"    ("{#(_)#}")
 | 
| 25507 | 100 | translations | 
| 101 |   "{#x, xs#}" == "{#x#} + {#xs#}"
 | |
| 102 |   "{#x#}" == "CONST single x"
 | |
| 103 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 104 | lemma count_empty [simp]: "count {#} a = 0"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 105 | by (simp add: Mempty_def in_multiset multiset_typedef) | 
| 10249 | 106 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 107 | lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 108 | by (simp add: single_def in_multiset multiset_typedef) | 
| 29901 | 109 | |
| 10249 | 110 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 111 | subsection {* Basic operations *}
 | 
| 10249 | 112 | |
| 113 | subsubsection {* Union *}
 | |
| 114 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 115 | lemma count_union [simp]: "count (M + N) a = count M a + count N a" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 116 | by (simp add: union_def in_multiset multiset_typedef) | 
| 10249 | 117 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 118 | instance multiset :: (type) cancel_comm_monoid_add proof | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 119 | qed (simp_all add: multiset_eq_iff) | 
| 10277 | 120 | |
| 10249 | 121 | |
| 122 | subsubsection {* Difference *}
 | |
| 123 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 124 | instantiation multiset :: (type) minus | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 125 | begin | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 126 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 127 | definition diff_def: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 128 | "M - N = Abs_multiset (\<lambda>a. count M a - count N a)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 129 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 130 | instance .. | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 131 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 132 | end | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 133 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 134 | lemma count_diff [simp]: "count (M - N) a = count M a - count N a" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 135 | by (simp add: diff_def in_multiset multiset_typedef) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 136 | |
| 17161 | 137 | lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 138 | by(simp add: multiset_eq_iff) | 
| 36903 | 139 | |
| 140 | lemma diff_cancel[simp]: "A - A = {#}"
 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 141 | by (rule multiset_eqI) simp | 
| 10249 | 142 | |
| 36903 | 143 | lemma diff_union_cancelR [simp]: "M + N - N = (M::'a multiset)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 144 | by(simp add: multiset_eq_iff) | 
| 10249 | 145 | |
| 36903 | 146 | lemma diff_union_cancelL [simp]: "N + M - N = (M::'a multiset)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 147 | by(simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 148 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 149 | lemma insert_DiffM: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 150 |   "x \<in># M \<Longrightarrow> {#x#} + (M - {#x#}) = M"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 151 | by (clarsimp simp: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 152 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 153 | lemma insert_DiffM2 [simp]: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 154 |   "x \<in># M \<Longrightarrow> M - {#x#} + {#x#} = M"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 155 | by (clarsimp simp: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 156 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 157 | lemma diff_right_commute: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 158 | "(M::'a multiset) - N - Q = M - Q - N" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 159 | by (auto simp add: multiset_eq_iff) | 
| 36903 | 160 | |
| 161 | lemma diff_add: | |
| 162 | "(M::'a multiset) - (N + Q) = M - N - Q" | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 163 | by (simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 164 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 165 | lemma diff_union_swap: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 166 |   "a \<noteq> b \<Longrightarrow> M - {#a#} + {#b#} = M + {#b#} - {#a#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 167 | by (auto simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 168 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 169 | lemma diff_union_single_conv: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 170 |   "a \<in># J \<Longrightarrow> I + J - {#a#} = I + (J - {#a#})"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 171 | by (simp add: multiset_eq_iff) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 172 | |
| 10249 | 173 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 174 | subsubsection {* Equality of multisets *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 175 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 176 | lemma single_not_empty [simp]: "{#a#} \<noteq> {#} \<and> {#} \<noteq> {#a#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 177 | by (simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 178 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 179 | lemma single_eq_single [simp]: "{#a#} = {#b#} \<longleftrightarrow> a = b"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 180 | by (auto simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 181 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 182 | lemma union_eq_empty [iff]: "M + N = {#} \<longleftrightarrow> M = {#} \<and> N = {#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 183 | by (auto simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 184 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 185 | lemma empty_eq_union [iff]: "{#} = M + N \<longleftrightarrow> M = {#} \<and> N = {#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 186 | by (auto simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 187 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 188 | lemma multi_self_add_other_not_self [simp]: "M = M + {#x#} \<longleftrightarrow> False"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 189 | by (auto simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 190 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 191 | lemma diff_single_trivial: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 192 |   "\<not> x \<in># M \<Longrightarrow> M - {#x#} = M"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 193 | by (auto simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 194 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 195 | lemma diff_single_eq_union: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 196 |   "x \<in># M \<Longrightarrow> M - {#x#} = N \<longleftrightarrow> M = N + {#x#}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 197 | by auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 198 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 199 | lemma union_single_eq_diff: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 200 |   "M + {#x#} = N \<Longrightarrow> M = N - {#x#}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 201 | by (auto dest: sym) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 202 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 203 | lemma union_single_eq_member: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 204 |   "M + {#x#} = N \<Longrightarrow> x \<in># N"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 205 | by auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 206 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 207 | lemma union_is_single: | 
| 36903 | 208 |   "M + N = {#a#} \<longleftrightarrow> M = {#a#} \<and> N={#} \<or> M = {#} \<and> N = {#a#}" (is "?lhs = ?rhs")proof
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 209 | assume ?rhs then show ?lhs by auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 210 | next | 
| 36903 | 211 | assume ?lhs thus ?rhs | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 212 | by(simp add: multiset_eq_iff split:if_splits) (metis add_is_1) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 213 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 214 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 215 | lemma single_is_union: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 216 |   "{#a#} = M + N \<longleftrightarrow> {#a#} = M \<and> N = {#} \<or> M = {#} \<and> {#a#} = N"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 217 |   by (auto simp add: eq_commute [of "{#a#}" "M + N"] union_is_single)
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 218 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 219 | lemma add_eq_conv_diff: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 220 |   "M + {#a#} = N + {#b#} \<longleftrightarrow> M = N \<and> a = b \<or> M = N - {#a#} + {#b#} \<and> N = M - {#b#} + {#a#}"  (is "?lhs = ?rhs")
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 221 | (* shorter: by (simp add: multiset_eq_iff) fastsimp *) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 222 | proof | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 223 | assume ?rhs then show ?lhs | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 224 |   by (auto simp add: add_assoc add_commute [of "{#b#}"])
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 225 | (drule sym, simp add: add_assoc [symmetric]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 226 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 227 | assume ?lhs | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 228 | show ?rhs | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 229 | proof (cases "a = b") | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 230 | case True with `?lhs` show ?thesis by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 231 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 232 | case False | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 233 |     from `?lhs` have "a \<in># N + {#b#}" by (rule union_single_eq_member)
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 234 | with False have "a \<in># N" by auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 235 |     moreover from `?lhs` have "M = N + {#b#} - {#a#}" by (rule union_single_eq_diff)
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 236 | moreover note False | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 237 |     ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"] diff_union_swap)
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 238 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 239 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 240 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 241 | lemma insert_noteq_member: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 242 |   assumes BC: "B + {#b#} = C + {#c#}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 243 | and bnotc: "b \<noteq> c" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 244 | shows "c \<in># B" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 245 | proof - | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 246 |   have "c \<in># C + {#c#}" by simp
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 247 |   have nc: "\<not> c \<in># {#b#}" using bnotc by simp
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 248 |   then have "c \<in># B + {#b#}" using BC by simp
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 249 | then show "c \<in># B" using nc by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 250 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 251 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 252 | lemma add_eq_conv_ex: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 253 |   "(M + {#a#} = N + {#b#}) =
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 254 |     (M = N \<and> a = b \<or> (\<exists>K. M = K + {#b#} \<and> N = K + {#a#}))"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 255 | by (auto simp add: add_eq_conv_diff) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 256 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 257 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 258 | subsubsection {* Pointwise ordering induced by count *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 259 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 260 | instantiation multiset :: (type) ordered_ab_semigroup_add_imp_le | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 261 | begin | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 262 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 263 | definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 264 | mset_le_def: "A \<le> B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 265 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 266 | definition less_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 267 | mset_less_def: "(A::'a multiset) < B \<longleftrightarrow> A \<le> B \<and> A \<noteq> B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 268 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 269 | instance proof | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 270 | qed (auto simp add: mset_le_def mset_less_def multiset_eq_iff intro: order_trans antisym) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 271 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 272 | end | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 273 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 274 | lemma mset_less_eqI: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 275 | "(\<And>x. count A x \<le> count B x) \<Longrightarrow> A \<le> B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 276 | by (simp add: mset_le_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 277 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 278 | lemma mset_le_exists_conv: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 279 | "(A::'a multiset) \<le> B \<longleftrightarrow> (\<exists>C. B = A + C)" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 280 | apply (unfold mset_le_def, rule iffI, rule_tac x = "B - A" in exI) | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 281 | apply (auto intro: multiset_eq_iff [THEN iffD2]) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 282 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 283 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 284 | lemma mset_le_mono_add_right_cancel [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 285 | "(A::'a multiset) + C \<le> B + C \<longleftrightarrow> A \<le> B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 286 | by (fact add_le_cancel_right) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 287 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 288 | lemma mset_le_mono_add_left_cancel [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 289 | "C + (A::'a multiset) \<le> C + B \<longleftrightarrow> A \<le> B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 290 | by (fact add_le_cancel_left) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 291 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 292 | lemma mset_le_mono_add: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 293 | "(A::'a multiset) \<le> B \<Longrightarrow> C \<le> D \<Longrightarrow> A + C \<le> B + D" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 294 | by (fact add_mono) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 295 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 296 | lemma mset_le_add_left [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 297 | "(A::'a multiset) \<le> A + B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 298 | unfolding mset_le_def by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 299 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 300 | lemma mset_le_add_right [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 301 | "B \<le> (A::'a multiset) + B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 302 | unfolding mset_le_def by auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 303 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 304 | lemma mset_le_single: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 305 |   "a :# B \<Longrightarrow> {#a#} \<le> B"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 306 | by (simp add: mset_le_def) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 307 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 308 | lemma multiset_diff_union_assoc: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 309 | "C \<le> B \<Longrightarrow> (A::'a multiset) + B - C = A + (B - C)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 310 | by (simp add: multiset_eq_iff mset_le_def) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 311 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 312 | lemma mset_le_multiset_union_diff_commute: | 
| 36867 | 313 | "B \<le> A \<Longrightarrow> (A::'a multiset) - B + C = A + C - B" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 314 | by (simp add: multiset_eq_iff mset_le_def) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 315 | |
| 39301 | 316 | lemma diff_le_self[simp]: "(M::'a multiset) - N \<le> M" | 
| 317 | by(simp add: mset_le_def) | |
| 318 | ||
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 319 | lemma mset_lessD: "A < B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 320 | apply (clarsimp simp: mset_le_def mset_less_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 321 | apply (erule_tac x=x in allE) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 322 | apply auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 323 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 324 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 325 | lemma mset_leD: "A \<le> B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 326 | apply (clarsimp simp: mset_le_def mset_less_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 327 | apply (erule_tac x = x in allE) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 328 | apply auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 329 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 330 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 331 | lemma mset_less_insertD: "(A + {#x#} < B) \<Longrightarrow> (x \<in># B \<and> A < B)"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 332 | apply (rule conjI) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 333 | apply (simp add: mset_lessD) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 334 | apply (clarsimp simp: mset_le_def mset_less_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 335 | apply safe | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 336 | apply (erule_tac x = a in allE) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 337 | apply (auto split: split_if_asm) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 338 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 339 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 340 | lemma mset_le_insertD: "(A + {#x#} \<le> B) \<Longrightarrow> (x \<in># B \<and> A \<le> B)"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 341 | apply (rule conjI) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 342 | apply (simp add: mset_leD) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 343 | apply (force simp: mset_le_def mset_less_def split: split_if_asm) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 344 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 345 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 346 | lemma mset_less_of_empty[simp]: "A < {#} \<longleftrightarrow> False"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 347 | by (auto simp add: mset_less_def mset_le_def multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 348 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 349 | lemma multi_psub_of_add_self[simp]: "A < A + {#x#}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 350 | by (auto simp: mset_le_def mset_less_def) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 351 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 352 | lemma multi_psub_self[simp]: "(A::'a multiset) < A = False" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 353 | by simp | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 354 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 355 | lemma mset_less_add_bothsides: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 356 |   "T + {#x#} < S + {#x#} \<Longrightarrow> T < S"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 357 | by (fact add_less_imp_less_right) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 358 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 359 | lemma mset_less_empty_nonempty: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 360 |   "{#} < S \<longleftrightarrow> S \<noteq> {#}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 361 | by (auto simp: mset_le_def mset_less_def) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 362 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 363 | lemma mset_less_diff_self: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 364 |   "c \<in># B \<Longrightarrow> B - {#c#} < B"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 365 | by (auto simp: mset_le_def mset_less_def multiset_eq_iff) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 366 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 367 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 368 | subsubsection {* Intersection *}
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 369 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 370 | instantiation multiset :: (type) semilattice_inf | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 371 | begin | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 372 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 373 | definition inf_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 374 | multiset_inter_def: "inf_multiset A B = A - (A - B)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 375 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 376 | instance proof - | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 377 | have aux: "\<And>m n q :: nat. m \<le> n \<Longrightarrow> m \<le> q \<Longrightarrow> m \<le> n - (n - q)" by arith | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 378 |   show "OFCLASS('a multiset, semilattice_inf_class)" proof
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 379 | qed (auto simp add: multiset_inter_def mset_le_def aux) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 380 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 381 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 382 | end | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 383 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 384 | abbreviation multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 385 | "multiset_inter \<equiv> inf" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 386 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 387 | lemma multiset_inter_count [simp]: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 388 | "count (A #\<inter> B) x = min (count A x) (count B x)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 389 | by (simp add: multiset_inter_def multiset_typedef) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 390 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 391 | lemma multiset_inter_single: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 392 | by (rule multiset_eqI) (auto simp add: multiset_inter_count) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 393 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 394 | lemma multiset_union_diff_commute: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 395 |   assumes "B #\<inter> C = {#}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 396 | shows "A + B - C = A - C + B" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 397 | proof (rule multiset_eqI) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 398 | fix x | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 399 | from assms have "min (count B x) (count C x) = 0" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 400 | by (auto simp add: multiset_inter_count multiset_eq_iff) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 401 | then have "count B x = 0 \<or> count C x = 0" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 402 | by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 403 | then show "count (A + B - C) x = count (A - C + B) x" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 404 | by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 405 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 406 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 407 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 408 | subsubsection {* Filter (with comprehension syntax) *}
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 409 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 410 | text {* Multiset comprehension *}
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 411 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 412 | definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 413 | "filter P M = Abs_multiset (\<lambda>x. if P x then count M x else 0)" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 414 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 415 | hide_const (open) filter | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 416 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 417 | lemma count_filter [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 418 | "count (Multiset.filter P M) a = (if P a then count M a else 0)" | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 419 | by (simp add: filter_def in_multiset multiset_typedef) | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 420 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 421 | lemma filter_empty [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 422 |   "Multiset.filter P {#} = {#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 423 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 424 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 425 | lemma filter_single [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 426 |   "Multiset.filter P {#x#} = (if P x then {#x#} else {#})"
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 427 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 428 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 429 | lemma filter_union [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 430 | "Multiset.filter P (M + N) = Multiset.filter P M + Multiset.filter P N" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 431 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 432 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 433 | lemma filter_diff [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 434 | "Multiset.filter P (M - N) = Multiset.filter P M - Multiset.filter P N" | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 435 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 436 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 437 | lemma filter_inter [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 438 | "Multiset.filter P (M #\<inter> N) = Multiset.filter P M #\<inter> Multiset.filter P N" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 439 | by (rule multiset_eqI) simp | 
| 10249 | 440 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 441 | syntax | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 442 |   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{# _ :# _./ _#})")
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 443 | syntax (xsymbol) | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 444 |   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{# _ \<in># _./ _#})")
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 445 | translations | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 446 |   "{#x \<in># M. P#}" == "CONST Multiset.filter (\<lambda>x. P) M"
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 447 | |
| 10249 | 448 | |
| 449 | subsubsection {* Set of elements *}
 | |
| 450 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 451 | definition set_of :: "'a multiset => 'a set" where | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 452 |   "set_of M = {x. x :# M}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 453 | |
| 17161 | 454 | lemma set_of_empty [simp]: "set_of {#} = {}"
 | 
| 26178 | 455 | by (simp add: set_of_def) | 
| 10249 | 456 | |
| 17161 | 457 | lemma set_of_single [simp]: "set_of {#b#} = {b}"
 | 
| 26178 | 458 | by (simp add: set_of_def) | 
| 10249 | 459 | |
| 17161 | 460 | lemma set_of_union [simp]: "set_of (M + N) = set_of M \<union> set_of N" | 
| 26178 | 461 | by (auto simp add: set_of_def) | 
| 10249 | 462 | |
| 17161 | 463 | lemma set_of_eq_empty_iff [simp]: "(set_of M = {}) = (M = {#})"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 464 | by (auto simp add: set_of_def multiset_eq_iff) | 
| 10249 | 465 | |
| 17161 | 466 | lemma mem_set_of_iff [simp]: "(x \<in> set_of M) = (x :# M)" | 
| 26178 | 467 | by (auto simp add: set_of_def) | 
| 26016 | 468 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 469 | lemma set_of_filter [simp]: "set_of {# x:#M. P x #} = set_of M \<inter> {x. P x}"
 | 
| 26178 | 470 | by (auto simp add: set_of_def) | 
| 10249 | 471 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 472 | lemma finite_set_of [iff]: "finite (set_of M)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 473 | using count [of M] by (simp add: multiset_def set_of_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 474 | |
| 10249 | 475 | |
| 476 | subsubsection {* Size *}
 | |
| 477 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 478 | instantiation multiset :: (type) size | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 479 | begin | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 480 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 481 | definition size_def: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 482 | "size M = setsum (count M) (set_of M)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 483 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 484 | instance .. | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 485 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 486 | end | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 487 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 488 | lemma size_empty [simp]: "size {#} = 0"
 | 
| 26178 | 489 | by (simp add: size_def) | 
| 10249 | 490 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 491 | lemma size_single [simp]: "size {#b#} = 1"
 | 
| 26178 | 492 | by (simp add: size_def) | 
| 10249 | 493 | |
| 17161 | 494 | lemma setsum_count_Int: | 
| 26178 | 495 | "finite A ==> setsum (count N) (A \<inter> set_of N) = setsum (count N) A" | 
| 496 | apply (induct rule: finite_induct) | |
| 497 | apply simp | |
| 498 | apply (simp add: Int_insert_left set_of_def) | |
| 499 | done | |
| 10249 | 500 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 501 | lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" | 
| 26178 | 502 | apply (unfold size_def) | 
| 503 | apply (subgoal_tac "count (M + N) = (\<lambda>a. count M a + count N a)") | |
| 504 | prefer 2 | |
| 505 | apply (rule ext, simp) | |
| 506 | apply (simp (no_asm_simp) add: setsum_Un_nat setsum_addf setsum_count_Int) | |
| 507 | apply (subst Int_commute) | |
| 508 | apply (simp (no_asm_simp) add: setsum_count_Int) | |
| 509 | done | |
| 10249 | 510 | |
| 17161 | 511 | lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 512 | by (auto simp add: size_def multiset_eq_iff) | 
| 26016 | 513 | |
| 514 | lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
 | |
| 26178 | 515 | by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) | 
| 10249 | 516 | |
| 17161 | 517 | lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \<exists>a. a :# M" | 
| 26178 | 518 | apply (unfold size_def) | 
| 519 | apply (drule setsum_SucD) | |
| 520 | apply auto | |
| 521 | done | |
| 10249 | 522 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 523 | lemma size_eq_Suc_imp_eq_union: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 524 | assumes "size M = Suc n" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 525 |   shows "\<exists>a N. M = N + {#a#}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 526 | proof - | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 527 | from assms obtain a where "a \<in># M" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 528 | by (erule size_eq_Suc_imp_elem [THEN exE]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 529 |   then have "M = M - {#a#} + {#a#}" by simp
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 530 | then show ?thesis by blast | 
| 23611 | 531 | qed | 
| 15869 | 532 | |
| 26016 | 533 | |
| 534 | subsection {* Induction and case splits *}
 | |
| 10249 | 535 | |
| 536 | lemma setsum_decr: | |
| 11701 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 wenzelm parents: 
11655diff
changeset | 537 | "finite F ==> (0::nat) < f a ==> | 
| 15072 | 538 | setsum (f (a := f a - 1)) F = (if a\<in>F then setsum f F - 1 else setsum f F)" | 
| 26178 | 539 | apply (induct rule: finite_induct) | 
| 540 | apply auto | |
| 541 | apply (drule_tac a = a in mk_disjoint_insert, auto) | |
| 542 | done | |
| 10249 | 543 | |
| 10313 | 544 | lemma rep_multiset_induct_aux: | 
| 26178 | 545 | assumes 1: "P (\<lambda>a. (0::nat))" | 
| 546 | and 2: "!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))" | |
| 547 | shows "\<forall>f. f \<in> multiset --> setsum f {x. f x \<noteq> 0} = n --> P f"
 | |
| 548 | apply (unfold multiset_def) | |
| 549 | apply (induct_tac n, simp, clarify) | |
| 550 | apply (subgoal_tac "f = (\<lambda>a.0)") | |
| 551 | apply simp | |
| 552 | apply (rule 1) | |
| 553 | apply (rule ext, force, clarify) | |
| 554 | apply (frule setsum_SucD, clarify) | |
| 555 | apply (rename_tac a) | |
| 556 | apply (subgoal_tac "finite {x. (f (a := f a - 1)) x > 0}")
 | |
| 557 | prefer 2 | |
| 558 | apply (rule finite_subset) | |
| 559 | prefer 2 | |
| 560 | apply assumption | |
| 561 | apply simp | |
| 562 | apply blast | |
| 563 | apply (subgoal_tac "f = (f (a := f a - 1))(a := (f (a := f a - 1)) a + 1)") | |
| 564 | prefer 2 | |
| 565 | apply (rule ext) | |
| 566 | apply (simp (no_asm_simp)) | |
| 567 | apply (erule ssubst, rule 2 [unfolded multiset_def], blast) | |
| 568 | apply (erule allE, erule impE, erule_tac [2] mp, blast) | |
| 569 | apply (simp (no_asm_simp) add: setsum_decr del: fun_upd_apply One_nat_def) | |
| 570 | apply (subgoal_tac "{x. x \<noteq> a --> f x \<noteq> 0} = {x. f x \<noteq> 0}")
 | |
| 571 | prefer 2 | |
| 572 | apply blast | |
| 573 | apply (subgoal_tac "{x. x \<noteq> a \<and> f x \<noteq> 0} = {x. f x \<noteq> 0} - {a}")
 | |
| 574 | prefer 2 | |
| 575 | apply blast | |
| 576 | apply (simp add: le_imp_diff_is_add setsum_diff1_nat cong: conj_cong) | |
| 577 | done | |
| 10249 | 578 | |
| 10313 | 579 | theorem rep_multiset_induct: | 
| 11464 | 580 | "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 | 581 | (!!f b. f \<in> multiset ==> P f ==> P (f (b := f b + 1))) ==> P f" | 
| 26178 | 582 | using rep_multiset_induct_aux by blast | 
| 10249 | 583 | |
| 18258 | 584 | theorem multiset_induct [case_names empty add, induct type: multiset]: | 
| 26178 | 585 | assumes empty: "P {#}"
 | 
| 586 |   and add: "!!M x. P M ==> P (M + {#x#})"
 | |
| 587 | shows "P M" | |
| 10249 | 588 | proof - | 
| 589 | note defns = union_def single_def Mempty_def | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 590 | note add' = add [unfolded defns, simplified] | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 591 | have aux: "\<And>a::'a. count (Abs_multiset (\<lambda>b. if b = a then 1 else 0)) = | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 592 | (\<lambda>b. if b = a then 1 else 0)" by (simp add: Abs_multiset_inverse in_multiset) | 
| 10249 | 593 | show ?thesis | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 594 | apply (rule count_inverse [THEN subst]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 595 | apply (rule count [THEN rep_multiset_induct]) | 
| 18258 | 596 | apply (rule empty [unfolded defns]) | 
| 15072 | 597 | apply (subgoal_tac "f(b := f b + 1) = (\<lambda>a. f a + (if a=b then 1 else 0))") | 
| 10249 | 598 | prefer 2 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 599 | apply (simp add: fun_eq_iff) | 
| 10249 | 600 | apply (erule ssubst) | 
| 17200 | 601 | apply (erule Abs_multiset_inverse [THEN subst]) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 602 | apply (drule add') | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 603 | apply (simp add: aux) | 
| 10249 | 604 | done | 
| 605 | qed | |
| 606 | ||
| 25610 | 607 | lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = A + {#a#}"
 | 
| 26178 | 608 | by (induct M) auto | 
| 25610 | 609 | |
| 610 | lemma multiset_cases [cases type, case_names empty add]: | |
| 26178 | 611 | assumes em:  "M = {#} \<Longrightarrow> P"
 | 
| 612 | assumes add: "\<And>N x. M = N + {#x#} \<Longrightarrow> P"
 | |
| 613 | shows "P" | |
| 25610 | 614 | proof (cases "M = {#}")
 | 
| 26145 | 615 |   assume "M = {#}" then show ?thesis using em by simp
 | 
| 25610 | 616 | next | 
| 617 |   assume "M \<noteq> {#}"
 | |
| 618 |   then obtain M' m where "M = M' + {#m#}" 
 | |
| 619 | by (blast dest: multi_nonempty_split) | |
| 26145 | 620 | then show ?thesis using add by simp | 
| 25610 | 621 | qed | 
| 622 | ||
| 623 | lemma multi_member_split: "x \<in># M \<Longrightarrow> \<exists>A. M = A + {#x#}"
 | |
| 26178 | 624 | apply (cases M) | 
| 625 | apply simp | |
| 626 | apply (rule_tac x="M - {#x#}" in exI, simp)
 | |
| 627 | done | |
| 25610 | 628 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 629 | lemma multi_drop_mem_not_eq: "c \<in># B \<Longrightarrow> B - {#c#} \<noteq> B"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 630 | by (cases "B = {#}") (auto dest: multi_member_split)
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 631 | |
| 26033 | 632 | lemma multiset_partition: "M = {# x:#M. P x #} + {# x:#M. \<not> P x #}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 633 | apply (subst multiset_eq_iff) | 
| 26178 | 634 | apply auto | 
| 635 | done | |
| 10249 | 636 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 637 | lemma mset_less_size: "(A::'a multiset) < B \<Longrightarrow> size A < size B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 638 | proof (induct A arbitrary: B) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 639 | case (empty M) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 640 |   then have "M \<noteq> {#}" by (simp add: mset_less_empty_nonempty)
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 641 |   then obtain M' x where "M = M' + {#x#}" 
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 642 | by (blast dest: multi_nonempty_split) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 643 | then show ?case by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 644 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 645 | case (add S x T) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 646 | have IH: "\<And>B. S < B \<Longrightarrow> size S < size B" by fact | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 647 |   have SxsubT: "S + {#x#} < T" by fact
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 648 | then have "x \<in># T" and "S < T" by (auto dest: mset_less_insertD) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 649 |   then obtain T' where T: "T = T' + {#x#}" 
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 650 | by (blast dest: multi_member_split) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 651 | then have "S < T'" using SxsubT | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 652 | by (blast intro: mset_less_add_bothsides) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 653 | then have "size S < size T'" using IH by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 654 | then show ?case using T by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 655 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 656 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 657 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 658 | subsubsection {* Strong induction and subset induction for multisets *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 659 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 660 | text {* Well-foundedness of proper subset operator: *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 661 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 662 | text {* proper multiset subset *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 663 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 664 | definition | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 665 |   mset_less_rel :: "('a multiset * 'a multiset) set" where
 | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 666 |   "mset_less_rel = {(A,B). A < B}"
 | 
| 10249 | 667 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 668 | lemma multiset_add_sub_el_shuffle: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 669 | assumes "c \<in># B" and "b \<noteq> c" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 670 |   shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 671 | proof - | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 672 |   from `c \<in># B` obtain A where B: "B = A + {#c#}" 
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 673 | by (blast dest: multi_member_split) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 674 |   have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 675 |   then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}" 
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 676 | by (simp add: add_ac) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 677 | then show ?thesis using B by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 678 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 679 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 680 | lemma wf_mset_less_rel: "wf mset_less_rel" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 681 | apply (unfold mset_less_rel_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 682 | apply (rule wf_measure [THEN wf_subset, where f1=size]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 683 | apply (clarsimp simp: measure_def inv_image_def mset_less_size) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 684 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 685 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 686 | text {* The induction rules: *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 687 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 688 | lemma full_multiset_induct [case_names less]: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 689 | assumes ih: "\<And>B. \<forall>(A::'a multiset). A < B \<longrightarrow> P A \<Longrightarrow> P B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 690 | shows "P B" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 691 | apply (rule wf_mset_less_rel [THEN wf_induct]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 692 | apply (rule ih, auto simp: mset_less_rel_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 693 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 694 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 695 | lemma multi_subset_induct [consumes 2, case_names empty add]: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 696 | assumes "F \<le> A" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 697 |   and empty: "P {#}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 698 |   and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (F + {#a#})"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 699 | shows "P F" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 700 | proof - | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 701 | from `F \<le> A` | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 702 | show ?thesis | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 703 | proof (induct F) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 704 |     show "P {#}" by fact
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 705 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 706 | fix x F | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 707 |     assume P: "F \<le> A \<Longrightarrow> P F" and i: "F + {#x#} \<le> A"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 708 |     show "P (F + {#x#})"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 709 | proof (rule insert) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 710 | from i show "x \<in># A" by (auto dest: mset_le_insertD) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 711 | from i have "F \<le> A" by (auto dest: mset_le_insertD) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 712 | with P show "P F" . | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 713 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 714 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 715 | qed | 
| 26145 | 716 | |
| 17161 | 717 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 718 | subsection {* Alternative representations *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 719 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 720 | subsubsection {* Lists *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 721 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 722 | primrec multiset_of :: "'a list \<Rightarrow> 'a multiset" where | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 723 |   "multiset_of [] = {#}" |
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 724 |   "multiset_of (a # x) = multiset_of x + {# a #}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 725 | |
| 37107 | 726 | lemma in_multiset_in_set: | 
| 727 | "x \<in># multiset_of xs \<longleftrightarrow> x \<in> set xs" | |
| 728 | by (induct xs) simp_all | |
| 729 | ||
| 730 | lemma count_multiset_of: | |
| 731 | "count (multiset_of xs) x = length (filter (\<lambda>y. x = y) xs)" | |
| 732 | by (induct xs) simp_all | |
| 733 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 734 | lemma multiset_of_zero_iff[simp]: "(multiset_of x = {#}) = (x = [])"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 735 | by (induct x) auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 736 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 737 | lemma multiset_of_zero_iff_right[simp]: "({#} = multiset_of x) = (x = [])"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 738 | by (induct x) auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 739 | |
| 40950 | 740 | lemma set_of_multiset_of[simp]: "set_of (multiset_of x) = set x" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 741 | by (induct x) auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 742 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 743 | lemma mem_set_multiset_eq: "x \<in> set xs = (x :# multiset_of xs)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 744 | by (induct xs) auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 745 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 746 | lemma multiset_of_append [simp]: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 747 | "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 748 | by (induct xs arbitrary: ys) (auto simp: add_ac) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 749 | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 750 | lemma multiset_of_filter: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 751 |   "multiset_of (filter P xs) = {#x :# multiset_of xs. P x #}"
 | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 752 | by (induct xs) simp_all | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 753 | |
| 40950 | 754 | lemma multiset_of_rev [simp]: | 
| 755 | "multiset_of (rev xs) = multiset_of xs" | |
| 756 | by (induct xs) simp_all | |
| 757 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 758 | lemma surj_multiset_of: "surj multiset_of" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 759 | apply (unfold surj_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 760 | apply (rule allI) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 761 | apply (rule_tac M = y in multiset_induct) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 762 | apply auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 763 | apply (rule_tac x = "x # xa" in exI) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 764 | apply auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 765 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 766 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 767 | lemma set_count_greater_0: "set x = {a. count (multiset_of x) a > 0}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 768 | by (induct x) auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 769 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 770 | lemma distinct_count_atmost_1: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 771 | "distinct x = (! a. count (multiset_of x) a = (if a \<in> set x then 1 else 0))" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 772 | apply (induct x, simp, rule iffI, simp_all) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 773 | apply (rule conjI) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 774 | apply (simp_all add: set_of_multiset_of [THEN sym] del: set_of_multiset_of) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 775 | apply (erule_tac x = a in allE, simp, clarify) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 776 | apply (erule_tac x = aa in allE, simp) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 777 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 778 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 779 | lemma multiset_of_eq_setD: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 780 | "multiset_of xs = multiset_of ys \<Longrightarrow> set xs = set ys" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 781 | by (rule) (auto simp add:multiset_eq_iff set_count_greater_0) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 782 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 783 | lemma set_eq_iff_multiset_of_eq_distinct: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 784 | "distinct x \<Longrightarrow> distinct y \<Longrightarrow> | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 785 | (set x = set y) = (multiset_of x = multiset_of y)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 786 | by (auto simp: multiset_eq_iff distinct_count_atmost_1) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 787 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 788 | lemma set_eq_iff_multiset_of_remdups_eq: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 789 | "(set x = set y) = (multiset_of (remdups x) = multiset_of (remdups y))" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 790 | apply (rule iffI) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 791 | apply (simp add: set_eq_iff_multiset_of_eq_distinct[THEN iffD1]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 792 | apply (drule distinct_remdups [THEN distinct_remdups | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 793 | [THEN set_eq_iff_multiset_of_eq_distinct [THEN iffD2]]]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 794 | apply simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 795 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 796 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 797 | lemma multiset_of_compl_union [simp]: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 798 | "multiset_of [x\<leftarrow>xs. P x] + multiset_of [x\<leftarrow>xs. \<not>P x] = multiset_of xs" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 799 | by (induct xs) (auto simp: add_ac) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 800 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 801 | lemma count_multiset_of_length_filter: | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 802 | "count (multiset_of xs) x = length (filter (\<lambda>y. x = y) xs)" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 803 | by (induct xs) auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 804 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 805 | lemma nth_mem_multiset_of: "i < length ls \<Longrightarrow> (ls ! i) :# multiset_of ls" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 806 | apply (induct ls arbitrary: i) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 807 | apply simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 808 | apply (case_tac i) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 809 | apply auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 810 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 811 | |
| 36903 | 812 | lemma multiset_of_remove1[simp]: | 
| 813 |   "multiset_of (remove1 a xs) = multiset_of xs - {#a#}"
 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 814 | by (induct xs) (auto simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 815 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 816 | lemma multiset_of_eq_length: | 
| 37107 | 817 | assumes "multiset_of xs = multiset_of ys" | 
| 818 | shows "length xs = length ys" | |
| 819 | using assms proof (induct xs arbitrary: ys) | |
| 820 | case Nil then show ?case by simp | |
| 821 | next | |
| 822 | case (Cons x xs) | |
| 823 | then have "x \<in># multiset_of ys" by (simp add: union_single_eq_member) | |
| 824 | then have "x \<in> set ys" by (simp add: in_multiset_in_set) | |
| 825 | from Cons.prems [symmetric] have "multiset_of xs = multiset_of (remove1 x ys)" | |
| 826 | by simp | |
| 827 | with Cons.hyps have "length xs = length (remove1 x ys)" . | |
| 828 | with `x \<in> set ys` show ?case | |
| 829 | by (auto simp add: length_remove1 dest: length_pos_if_in_set) | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 830 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 831 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 832 | lemma multiset_of_eq_length_filter: | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 833 | assumes "multiset_of xs = multiset_of ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 834 | shows "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) ys)" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 835 | proof (cases "z \<in># multiset_of xs") | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 836 | case False | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 837 | moreover have "\<not> z \<in># multiset_of ys" using assms False by simp | 
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 838 | ultimately show ?thesis by (simp add: count_multiset_of_length_filter) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 839 | next | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 840 | case True | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 841 | moreover have "z \<in># multiset_of ys" using assms True by simp | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 842 | show ?thesis using assms proof (induct xs arbitrary: ys) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 843 | case Nil then show ?case by simp | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 844 | next | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 845 | case (Cons x xs) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 846 | from `multiset_of (x # xs) = multiset_of ys` [symmetric] | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 847 | have *: "multiset_of xs = multiset_of (remove1 x ys)" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 848 | and "x \<in> set ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 849 | by (auto simp add: mem_set_multiset_eq) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 850 | from * have "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) (remove1 x ys))" by (rule Cons.hyps) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 851 | moreover from `x \<in> set ys` have "length (filter (\<lambda>y. x = y) ys) > 0" by (simp add: filter_empty_conv) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 852 | ultimately show ?case using `x \<in> set ys` | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 853 | by (simp add: filter_remove1) (auto simp add: length_remove1) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 854 | qed | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 855 | qed | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 856 | |
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 857 | context linorder | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 858 | begin | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 859 | |
| 40210 
aee7ef725330
sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
 haftmann parents: 
39533diff
changeset | 860 | lemma multiset_of_insort [simp]: | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 861 |   "multiset_of (insort_key k x xs) = {#x#} + multiset_of xs"
 | 
| 37107 | 862 | by (induct xs) (simp_all add: ac_simps) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 863 | |
| 40210 
aee7ef725330
sorting: avoid _key suffix if lemma applies both to simple and generalized variant; generalized insort_insert to insort_insert_key; additional lemmas
 haftmann parents: 
39533diff
changeset | 864 | lemma multiset_of_sort [simp]: | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 865 | "multiset_of (sort_key k xs) = multiset_of xs" | 
| 37107 | 866 | by (induct xs) (simp_all add: ac_simps) | 
| 867 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 868 | text {*
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 869 | This lemma shows which properties suffice to show that a function | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 870 |   @{text "f"} with @{text "f xs = ys"} behaves like sort.
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 871 | *} | 
| 37074 | 872 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 873 | lemma properties_for_sort_key: | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 874 | assumes "multiset_of ys = multiset_of xs" | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 875 | and "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>x. f k = f x) ys = filter (\<lambda>x. f k = f x) xs" | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 876 | and "sorted (map f ys)" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 877 | shows "sort_key f xs = ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 878 | using assms proof (induct xs arbitrary: ys) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 879 | case Nil then show ?case by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 880 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 881 | case (Cons x xs) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 882 | from Cons.prems(2) have | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 883 | "\<forall>k \<in> set ys. filter (\<lambda>x. f k = f x) (remove1 x ys) = filter (\<lambda>x. f k = f x) xs" | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 884 | by (simp add: filter_remove1) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 885 | with Cons.prems have "sort_key f xs = remove1 x ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 886 | by (auto intro!: Cons.hyps simp add: sorted_map_remove1) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 887 | moreover from Cons.prems have "x \<in> set ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 888 | by (auto simp add: mem_set_multiset_eq intro!: ccontr) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 889 | ultimately show ?case using Cons.prems by (simp add: insort_key_remove1) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 890 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 891 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 892 | lemma properties_for_sort: | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 893 | assumes multiset: "multiset_of ys = multiset_of xs" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 894 | and "sorted ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 895 | shows "sort xs = ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 896 | proof (rule properties_for_sort_key) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 897 | from multiset show "multiset_of ys = multiset_of xs" . | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 898 | from `sorted ys` show "sorted (map (\<lambda>x. x) ys)" by simp | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 899 | from multiset have "\<And>k. length (filter (\<lambda>y. k = y) ys) = length (filter (\<lambda>x. k = x) xs)" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 900 | by (rule multiset_of_eq_length_filter) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 901 | then have "\<And>k. replicate (length (filter (\<lambda>y. k = y) ys)) k = replicate (length (filter (\<lambda>x. k = x) xs)) k" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 902 | by simp | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 903 | then show "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>y. k = y) ys = filter (\<lambda>x. k = x) xs" | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 904 | by (simp add: replicate_length_filter) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 905 | qed | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 906 | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 907 | lemma sort_key_by_quicksort: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 908 | "sort_key f xs = sort_key f [x\<leftarrow>xs. f x < f (xs ! (length xs div 2))] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 909 | @ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 910 | @ sort_key f [x\<leftarrow>xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs") | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 911 | proof (rule properties_for_sort_key) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 912 | show "multiset_of ?rhs = multiset_of ?lhs" | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 913 | by (rule multiset_eqI) (auto simp add: multiset_of_filter) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 914 | next | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 915 | show "sorted (map f ?rhs)" | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 916 | by (auto simp add: sorted_append intro: sorted_map_same) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 917 | next | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 918 | fix l | 
| 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 919 | assume "l \<in> set ?rhs" | 
| 40346 | 920 | let ?pivot = "f (xs ! (length xs div 2))" | 
| 921 | have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto | |
| 40306 | 922 | have "[x \<leftarrow> sort_key f xs . f x = f l] = [x \<leftarrow> xs. f x = f l]" | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 923 | unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) | 
| 40346 | 924 | with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp | 
| 925 | have "\<And>x P. P (f x) ?pivot \<and> f l = f x \<longleftrightarrow> P (f l) ?pivot \<and> f l = f x" by auto | |
| 926 | then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] = | |
| 927 | [x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp | |
| 928 | note *** = this [of "op <"] this [of "op >"] this [of "op ="] | |
| 40306 | 929 | show "[x \<leftarrow> ?rhs. f l = f x] = [x \<leftarrow> ?lhs. f l = f x]" | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 930 | proof (cases "f l" ?pivot rule: linorder_cases) | 
| 40307 | 931 | case less then moreover have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto | 
| 932 | ultimately show ?thesis | |
| 40346 | 933 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 934 | next | 
| 40306 | 935 | case equal then show ?thesis | 
| 40346 | 936 | by (simp add: * less_le) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 937 | next | 
| 40307 | 938 | case greater then moreover have "f l \<noteq> ?pivot" and "\<not> f l < ?pivot" by auto | 
| 939 | ultimately show ?thesis | |
| 40346 | 940 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40306 | 941 | qed | 
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 942 | qed | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 943 | |
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 944 | lemma sort_by_quicksort: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 945 | "sort xs = sort [x\<leftarrow>xs. x < xs ! (length xs div 2)] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 946 | @ [x\<leftarrow>xs. x = xs ! (length xs div 2)] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 947 | @ sort [x\<leftarrow>xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 948 | using sort_key_by_quicksort [of "\<lambda>x. x", symmetric] by simp | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 949 | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 950 | text {* A stable parametrized quicksort *}
 | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 951 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 952 | definition part :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'b list \<times> 'b list \<times> 'b list" where
 | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 953 | "part f pivot xs = ([x \<leftarrow> xs. f x < pivot], [x \<leftarrow> xs. f x = pivot], [x \<leftarrow> xs. pivot < f x])" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 954 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 955 | lemma part_code [code]: | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 956 | "part f pivot [] = ([], [], [])" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 957 | "part f pivot (x # xs) = (let (lts, eqs, gts) = part f pivot xs; x' = f x in | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 958 | if x' < pivot then (x # lts, eqs, gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 959 | else if x' > pivot then (lts, eqs, x # gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 960 | else (lts, x # eqs, gts))" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 961 | by (auto simp add: part_def Let_def split_def) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 962 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 963 | lemma sort_key_by_quicksort_code [code]: | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 964 | "sort_key f xs = (case xs of [] \<Rightarrow> [] | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 965 | | [x] \<Rightarrow> xs | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 966 | | [x, y] \<Rightarrow> (if f x \<le> f y then xs else [y, x]) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 967 | | _ \<Rightarrow> (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 968 | in sort_key f lts @ eqs @ sort_key f gts))" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 969 | proof (cases xs) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 970 | case Nil then show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 971 | next | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 972 | case (Cons _ ys) note hyps = Cons show ?thesis proof (cases ys) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 973 | case Nil with hyps show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 974 | next | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 975 | case (Cons _ zs) note hyps = hyps Cons show ?thesis proof (cases zs) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 976 | case Nil with hyps show ?thesis by auto | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 977 | next | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 978 | case Cons | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 979 | from sort_key_by_quicksort [of f xs] | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 980 | have "sort_key f xs = (let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 981 | in sort_key f lts @ eqs @ sort_key f gts)" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 982 | by (simp only: split_def Let_def part_def fst_conv snd_conv) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 983 | with hyps Cons show ?thesis by (simp only: list.cases) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 984 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 985 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 986 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 987 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 988 | end | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 989 | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 990 | hide_const (open) part | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 991 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 992 | lemma multiset_of_remdups_le: "multiset_of (remdups xs) \<le> multiset_of xs" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 993 | by (induct xs) (auto intro: order_trans) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 994 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 995 | lemma multiset_of_update: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 996 |   "i < length ls \<Longrightarrow> multiset_of (ls[i := v]) = multiset_of ls - {#ls ! i#} + {#v#}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 997 | proof (induct ls arbitrary: i) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 998 | case Nil then show ?case by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 999 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1000 | case (Cons x xs) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1001 | show ?case | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1002 | proof (cases i) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1003 | case 0 then show ?thesis by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1004 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1005 | case (Suc i') | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1006 | with Cons show ?thesis | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1007 | apply simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1008 | apply (subst add_assoc) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1009 |       apply (subst add_commute [of "{#v#}" "{#x#}"])
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1010 | apply (subst add_assoc [symmetric]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1011 | apply simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1012 | apply (rule mset_le_multiset_union_diff_commute) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1013 | apply (simp add: mset_le_single nth_mem_multiset_of) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1014 | done | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1015 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1016 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1017 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1018 | lemma multiset_of_swap: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1019 | "i < length ls \<Longrightarrow> j < length ls \<Longrightarrow> | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1020 | multiset_of (ls[j := ls ! i, i := ls ! j]) = multiset_of ls" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1021 | by (cases "i = j") (simp_all add: multiset_of_update nth_mem_multiset_of) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1022 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1023 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1024 | subsubsection {* Association lists -- including rudimentary code generation *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1025 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1026 | definition count_of :: "('a \<times> nat) list \<Rightarrow> 'a \<Rightarrow> nat" where
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1027 | "count_of xs x = (case map_of xs x of None \<Rightarrow> 0 | Some n \<Rightarrow> n)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1028 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1029 | lemma count_of_multiset: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1030 | "count_of xs \<in> multiset" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1031 | proof - | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1032 |   let ?A = "{x::'a. 0 < (case map_of xs x of None \<Rightarrow> 0\<Colon>nat | Some (n\<Colon>nat) \<Rightarrow> n)}"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1033 | have "?A \<subseteq> dom (map_of xs)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1034 | proof | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1035 | fix x | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1036 | assume "x \<in> ?A" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1037 | then have "0 < (case map_of xs x of None \<Rightarrow> 0\<Colon>nat | Some (n\<Colon>nat) \<Rightarrow> n)" by simp | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1038 | then have "map_of xs x \<noteq> None" by (cases "map_of xs x") auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1039 | then show "x \<in> dom (map_of xs)" by auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1040 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1041 | with finite_dom_map_of [of xs] have "finite ?A" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1042 | by (auto intro: finite_subset) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1043 | then show ?thesis | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1044 | by (simp add: count_of_def fun_eq_iff multiset_def) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1045 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1046 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1047 | lemma count_simps [simp]: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1048 | "count_of [] = (\<lambda>_. 0)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1049 | "count_of ((x, n) # xs) = (\<lambda>y. if x = y then n else count_of xs y)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1050 | by (simp_all add: count_of_def fun_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1051 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1052 | lemma count_of_empty: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1053 | "x \<notin> fst ` set xs \<Longrightarrow> count_of xs x = 0" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1054 | by (induct xs) (simp_all add: count_of_def) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1055 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1056 | lemma count_of_filter: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1057 | "count_of (filter (P \<circ> fst) xs) x = (if P x then count_of xs x else 0)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1058 | by (induct xs) auto | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1059 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1060 | definition Bag :: "('a \<times> nat) list \<Rightarrow> 'a multiset" where
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1061 | "Bag xs = Abs_multiset (count_of xs)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1062 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1063 | code_datatype Bag | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1064 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1065 | lemma count_Bag [simp, code]: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1066 | "count (Bag xs) = count_of xs" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1067 | by (simp add: Bag_def count_of_multiset Abs_multiset_inverse) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1068 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1069 | lemma Mempty_Bag [code]: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1070 |   "{#} = Bag []"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1071 | by (simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1072 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1073 | lemma single_Bag [code]: | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1074 |   "{#x#} = Bag [(x, 1)]"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1075 | by (simp add: multiset_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1076 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 1077 | lemma filter_Bag [code]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 1078 | "Multiset.filter P (Bag xs) = Bag (filter (P \<circ> fst) xs)" | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 1079 | by (rule multiset_eqI) (simp add: count_of_filter) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1080 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1081 | lemma mset_less_eq_Bag [code]: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1082 | "Bag xs \<le> A \<longleftrightarrow> (\<forall>(x, n) \<in> set xs. count_of xs x \<le> count A x)" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1083 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1084 | proof | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1085 | assume ?lhs then show ?rhs | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1086 | by (auto simp add: mset_le_def count_Bag) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1087 | next | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1088 | assume ?rhs | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1089 | show ?lhs | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1090 | proof (rule mset_less_eqI) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1091 | fix x | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1092 | from `?rhs` have "count_of xs x \<le> count A x" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1093 | by (cases "x \<in> fst ` set xs") (auto simp add: count_of_empty) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1094 | then show "count (Bag xs) x \<le> count A x" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1095 | by (simp add: mset_le_def count_Bag) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1096 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1097 | qed | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1098 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38287diff
changeset | 1099 | instantiation multiset :: (equal) equal | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1100 | begin | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1101 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1102 | definition | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38287diff
changeset | 1103 | "HOL.equal A B \<longleftrightarrow> (A::'a multiset) \<le> B \<and> B \<le> A" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1104 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1105 | instance proof | 
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38287diff
changeset | 1106 | qed (simp add: equal_multiset_def eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1107 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1108 | end | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1109 | |
| 38857 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38287diff
changeset | 1110 | lemma [code nbe]: | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38287diff
changeset | 1111 | "HOL.equal (A :: 'a::equal multiset) A \<longleftrightarrow> True" | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38287diff
changeset | 1112 | by (fact equal_refl) | 
| 
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
 haftmann parents: 
38287diff
changeset | 1113 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1114 | definition (in term_syntax) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1115 |   bagify :: "('a\<Colon>typerep \<times> nat) list \<times> (unit \<Rightarrow> Code_Evaluation.term)
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1116 | \<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1117 |   [code_unfold]: "bagify xs = Code_Evaluation.valtermify Bag {\<cdot>} xs"
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1118 | |
| 37751 | 1119 | notation fcomp (infixl "\<circ>>" 60) | 
| 1120 | notation scomp (infixl "\<circ>\<rightarrow>" 60) | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1121 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1122 | instantiation multiset :: (random) random | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1123 | begin | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1124 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1125 | definition | 
| 37751 | 1126 | "Quickcheck.random i = Quickcheck.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (bagify xs))" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1127 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1128 | instance .. | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1129 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1130 | end | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1131 | |
| 37751 | 1132 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 1133 | no_notation scomp (infixl "\<circ>\<rightarrow>" 60) | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1134 | |
| 36176 
3fe7e97ccca8
replaced generic 'hide' command by more conventional 'hide_class', 'hide_type', 'hide_const', 'hide_fact' -- frees some popular keywords;
 wenzelm parents: 
35712diff
changeset | 1135 | hide_const (open) bagify | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1136 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1137 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1138 | subsection {* The multiset order *}
 | 
| 10249 | 1139 | |
| 1140 | subsubsection {* Well-foundedness *}
 | |
| 1141 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1142 | definition mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
 | 
| 37765 | 1143 |   "mult1 r = {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
 | 
| 23751 | 1144 | (\<forall>b. b :# K --> (b, a) \<in> r)}" | 
| 10249 | 1145 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1146 | definition mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
 | 
| 37765 | 1147 | "mult r = (mult1 r)\<^sup>+" | 
| 10249 | 1148 | |
| 23751 | 1149 | lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
 | 
| 26178 | 1150 | by (simp add: mult1_def) | 
| 10249 | 1151 | |
| 23751 | 1152 | lemma less_add: "(N, M0 + {#a#}) \<in> mult1 r ==>
 | 
| 1153 |     (\<exists>M. (M, M0) \<in> mult1 r \<and> N = M + {#a#}) \<or>
 | |
| 1154 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K)" | |
| 19582 | 1155 | (is "_ \<Longrightarrow> ?case1 (mult1 r) \<or> ?case2") | 
| 10249 | 1156 | proof (unfold mult1_def) | 
| 23751 | 1157 | let ?r = "\<lambda>K a. \<forall>b. b :# K --> (b, a) \<in> r" | 
| 11464 | 1158 |   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
 | 
| 23751 | 1159 |   let ?case1 = "?case1 {(N, M). ?R N M}"
 | 
| 10249 | 1160 | |
| 23751 | 1161 |   assume "(N, M0 + {#a#}) \<in> {(N, M). ?R N M}"
 | 
| 18258 | 1162 | then have "\<exists>a' M0' K. | 
| 11464 | 1163 |       M0 + {#a#} = M0' + {#a'#} \<and> N = M0' + K \<and> ?r K a'" by simp
 | 
| 18258 | 1164 | then show "?case1 \<or> ?case2" | 
| 10249 | 1165 | proof (elim exE conjE) | 
| 1166 | fix a' M0' K | |
| 1167 | assume N: "N = M0' + K" and r: "?r K a'" | |
| 1168 |     assume "M0 + {#a#} = M0' + {#a'#}"
 | |
| 18258 | 1169 | then have "M0 = M0' \<and> a = a' \<or> | 
| 11464 | 1170 |         (\<exists>K'. M0 = K' + {#a'#} \<and> M0' = K' + {#a#})"
 | 
| 10249 | 1171 | by (simp only: add_eq_conv_ex) | 
| 18258 | 1172 | then show ?thesis | 
| 10249 | 1173 | proof (elim disjE conjE exE) | 
| 1174 | assume "M0 = M0'" "a = a'" | |
| 11464 | 1175 | with N r have "?r K a \<and> N = M0 + K" by simp | 
| 18258 | 1176 | then have ?case2 .. then show ?thesis .. | 
| 10249 | 1177 | next | 
| 1178 | fix K' | |
| 1179 |       assume "M0' = K' + {#a#}"
 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1180 |       with N have n: "N = K' + K + {#a#}" by (simp add: add_ac)
 | 
| 10249 | 1181 | |
| 1182 |       assume "M0 = K' + {#a'#}"
 | |
| 1183 | with r have "?R (K' + K) M0" by blast | |
| 18258 | 1184 | with n have ?case1 by simp then show ?thesis .. | 
| 10249 | 1185 | qed | 
| 1186 | qed | |
| 1187 | qed | |
| 1188 | ||
| 23751 | 1189 | lemma all_accessible: "wf r ==> \<forall>M. M \<in> acc (mult1 r)" | 
| 10249 | 1190 | proof | 
| 1191 | let ?R = "mult1 r" | |
| 1192 | let ?W = "acc ?R" | |
| 1193 |   {
 | |
| 1194 | fix M M0 a | |
| 23751 | 1195 | assume M0: "M0 \<in> ?W" | 
| 1196 |       and wf_hyp: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | |
| 1197 |       and acc_hyp: "\<forall>M. (M, M0) \<in> ?R --> M + {#a#} \<in> ?W"
 | |
| 1198 |     have "M0 + {#a#} \<in> ?W"
 | |
| 1199 |     proof (rule accI [of "M0 + {#a#}"])
 | |
| 10249 | 1200 | fix N | 
| 23751 | 1201 |       assume "(N, M0 + {#a#}) \<in> ?R"
 | 
| 1202 |       then have "((\<exists>M. (M, M0) \<in> ?R \<and> N = M + {#a#}) \<or>
 | |
| 1203 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K))" | |
| 10249 | 1204 | by (rule less_add) | 
| 23751 | 1205 | then show "N \<in> ?W" | 
| 10249 | 1206 | proof (elim exE disjE conjE) | 
| 23751 | 1207 |         fix M assume "(M, M0) \<in> ?R" and N: "N = M + {#a#}"
 | 
| 1208 |         from acc_hyp have "(M, M0) \<in> ?R --> M + {#a#} \<in> ?W" ..
 | |
| 1209 |         from this and `(M, M0) \<in> ?R` have "M + {#a#} \<in> ?W" ..
 | |
| 1210 | then show "N \<in> ?W" by (simp only: N) | |
| 10249 | 1211 | next | 
| 1212 | fix K | |
| 1213 | assume N: "N = M0 + K" | |
| 23751 | 1214 | assume "\<forall>b. b :# K --> (b, a) \<in> r" | 
| 1215 | then have "M0 + K \<in> ?W" | |
| 10249 | 1216 | proof (induct K) | 
| 18730 | 1217 | case empty | 
| 23751 | 1218 |           from M0 show "M0 + {#} \<in> ?W" by simp
 | 
| 18730 | 1219 | next | 
| 1220 | case (add K x) | |
| 23751 | 1221 | from add.prems have "(x, a) \<in> r" by simp | 
| 1222 |           with wf_hyp have "\<forall>M \<in> ?W. M + {#x#} \<in> ?W" by blast
 | |
| 1223 | moreover from add have "M0 + K \<in> ?W" by simp | |
| 1224 |           ultimately have "(M0 + K) + {#x#} \<in> ?W" ..
 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1225 |           then show "M0 + (K + {#x#}) \<in> ?W" by (simp only: add_assoc)
 | 
| 10249 | 1226 | qed | 
| 23751 | 1227 | then show "N \<in> ?W" by (simp only: N) | 
| 10249 | 1228 | qed | 
| 1229 | qed | |
| 1230 | } note tedious_reasoning = this | |
| 1231 | ||
| 23751 | 1232 | assume wf: "wf r" | 
| 10249 | 1233 | fix M | 
| 23751 | 1234 | show "M \<in> ?W" | 
| 10249 | 1235 | proof (induct M) | 
| 23751 | 1236 |     show "{#} \<in> ?W"
 | 
| 10249 | 1237 | proof (rule accI) | 
| 23751 | 1238 |       fix b assume "(b, {#}) \<in> ?R"
 | 
| 1239 | with not_less_empty show "b \<in> ?W" by contradiction | |
| 10249 | 1240 | qed | 
| 1241 | ||
| 23751 | 1242 | fix M a assume "M \<in> ?W" | 
| 1243 |     from wf have "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | |
| 10249 | 1244 | proof induct | 
| 1245 | fix a | |
| 23751 | 1246 |       assume r: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | 
| 1247 |       show "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | |
| 10249 | 1248 | proof | 
| 23751 | 1249 | fix M assume "M \<in> ?W" | 
| 1250 |         then show "M + {#a#} \<in> ?W"
 | |
| 23373 | 1251 | by (rule acc_induct) (rule tedious_reasoning [OF _ r]) | 
| 10249 | 1252 | qed | 
| 1253 | qed | |
| 23751 | 1254 |     from this and `M \<in> ?W` show "M + {#a#} \<in> ?W" ..
 | 
| 10249 | 1255 | qed | 
| 1256 | qed | |
| 1257 | ||
| 23751 | 1258 | theorem wf_mult1: "wf r ==> wf (mult1 r)" | 
| 26178 | 1259 | by (rule acc_wfI) (rule all_accessible) | 
| 10249 | 1260 | |
| 23751 | 1261 | theorem wf_mult: "wf r ==> wf (mult r)" | 
| 26178 | 1262 | unfolding mult_def by (rule wf_trancl) (rule wf_mult1) | 
| 10249 | 1263 | |
| 1264 | ||
| 1265 | subsubsection {* Closure-free presentation *}
 | |
| 1266 | ||
| 1267 | text {* One direction. *}
 | |
| 1268 | ||
| 1269 | lemma mult_implies_one_step: | |
| 23751 | 1270 | "trans r ==> (M, N) \<in> mult r ==> | 
| 11464 | 1271 |     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
 | 
| 23751 | 1272 | (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r)" | 
| 26178 | 1273 | apply (unfold mult_def mult1_def set_of_def) | 
| 1274 | apply (erule converse_trancl_induct, clarify) | |
| 1275 | apply (rule_tac x = M0 in exI, simp, clarify) | |
| 1276 | apply (case_tac "a :# K") | |
| 1277 | apply (rule_tac x = I in exI) | |
| 1278 | apply (simp (no_asm)) | |
| 1279 |  apply (rule_tac x = "(K - {#a#}) + Ka" in exI)
 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1280 | apply (simp (no_asm_simp) add: add_assoc [symmetric]) | 
| 26178 | 1281 |  apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong)
 | 
| 1282 | apply (simp add: diff_union_single_conv) | |
| 1283 | apply (simp (no_asm_use) add: trans_def) | |
| 1284 | apply blast | |
| 1285 | apply (subgoal_tac "a :# I") | |
| 1286 |  apply (rule_tac x = "I - {#a#}" in exI)
 | |
| 1287 |  apply (rule_tac x = "J + {#a#}" in exI)
 | |
| 1288 | apply (rule_tac x = "K + Ka" in exI) | |
| 1289 | apply (rule conjI) | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1290 | apply (simp add: multiset_eq_iff split: nat_diff_split) | 
| 26178 | 1291 | apply (rule conjI) | 
| 1292 |   apply (drule_tac f = "\<lambda>M. M - {#a#}" in arg_cong, simp)
 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1293 | apply (simp add: multiset_eq_iff split: nat_diff_split) | 
| 26178 | 1294 | apply (simp (no_asm_use) add: trans_def) | 
| 1295 | apply blast | |
| 1296 | apply (subgoal_tac "a :# (M0 + {#a#})")
 | |
| 1297 | apply simp | |
| 1298 | apply (simp (no_asm)) | |
| 1299 | done | |
| 10249 | 1300 | |
| 1301 | lemma one_step_implies_mult_aux: | |
| 23751 | 1302 | "trans r ==> | 
| 1303 |     \<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))
 | |
| 1304 | --> (I + K, I + J) \<in> mult r" | |
| 26178 | 1305 | apply (induct_tac n, auto) | 
| 1306 | apply (frule size_eq_Suc_imp_eq_union, clarify) | |
| 1307 | apply (rename_tac "J'", simp) | |
| 1308 | apply (erule notE, auto) | |
| 1309 | apply (case_tac "J' = {#}")
 | |
| 1310 | apply (simp add: mult_def) | |
| 1311 | apply (rule r_into_trancl) | |
| 1312 | apply (simp add: mult1_def set_of_def, blast) | |
| 1313 | txt {* Now we know @{term "J' \<noteq> {#}"}. *}
 | |
| 1314 | apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition) | |
| 1315 | apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp) | |
| 1316 | apply (erule ssubst) | |
| 1317 | apply (simp add: Ball_def, auto) | |
| 1318 | apply (subgoal_tac | |
| 1319 |   "((I + {# x :# K. (x, a) \<in> r #}) + {# x :# K. (x, a) \<notin> r #},
 | |
| 1320 |     (I + {# x :# K. (x, a) \<in> r #}) + J') \<in> mult r")
 | |
| 1321 | prefer 2 | |
| 1322 | apply force | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1323 | apply (simp (no_asm_use) add: add_assoc [symmetric] mult_def) | 
| 26178 | 1324 | apply (erule trancl_trans) | 
| 1325 | apply (rule r_into_trancl) | |
| 1326 | apply (simp add: mult1_def set_of_def) | |
| 1327 | apply (rule_tac x = a in exI) | |
| 1328 | apply (rule_tac x = "I + J'" in exI) | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1329 | apply (simp add: add_ac) | 
| 26178 | 1330 | done | 
| 10249 | 1331 | |
| 17161 | 1332 | lemma one_step_implies_mult: | 
| 23751 | 1333 |   "trans r ==> J \<noteq> {#} ==> \<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r
 | 
| 1334 | ==> (I + K, I + J) \<in> mult r" | |
| 26178 | 1335 | using one_step_implies_mult_aux by blast | 
| 10249 | 1336 | |
| 1337 | ||
| 1338 | subsubsection {* Partial-order properties *}
 | |
| 1339 | ||
| 35273 | 1340 | definition less_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<#" 50) where | 
| 1341 |   "M' <# M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
 | |
| 10249 | 1342 | |
| 35273 | 1343 | definition le_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<=#" 50) where | 
| 1344 | "M' <=# M \<longleftrightarrow> M' <# M \<or> M' = M" | |
| 1345 | ||
| 35308 | 1346 | notation (xsymbols) less_multiset (infix "\<subset>#" 50) | 
| 1347 | notation (xsymbols) le_multiset (infix "\<subseteq>#" 50) | |
| 10249 | 1348 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1349 | interpretation multiset_order: order le_multiset less_multiset | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1350 | proof - | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1351 | have irrefl: "\<And>M :: 'a multiset. \<not> M \<subset># M" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1352 | proof | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1353 | fix M :: "'a multiset" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1354 | assume "M \<subset># M" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1355 |     then have MM: "(M, M) \<in> mult {(x, y). x < y}" by (simp add: less_multiset_def)
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1356 |     have "trans {(x'::'a, x). x' < x}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1357 | by (rule transI) simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1358 | moreover note MM | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1359 | ultimately have "\<exists>I J K. M = I + J \<and> M = I + K | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1360 |       \<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1361 | by (rule mult_implies_one_step) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1362 | then obtain I J K where "M = I + J" and "M = I + K" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1363 |       and "J \<noteq> {#}" and "(\<forall>k\<in>set_of K. \<exists>j\<in>set_of J. (k, j) \<in> {(x, y). x < y})" by blast
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1364 |     then have aux1: "K \<noteq> {#}" and aux2: "\<forall>k\<in>set_of K. \<exists>j\<in>set_of K. k < j" by auto
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1365 | have "finite (set_of K)" by simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1366 | moreover note aux2 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1367 |     ultimately have "set_of K = {}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1368 | by (induct rule: finite_induct) (auto intro: order_less_trans) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1369 | with aux1 show False by simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1370 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1371 | have trans: "\<And>K M N :: 'a multiset. K \<subset># M \<Longrightarrow> M \<subset># N \<Longrightarrow> K \<subset># N" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1372 | unfolding less_multiset_def mult_def by (blast intro: trancl_trans) | 
| 36635 
080b755377c0
locale predicates of classes carry a mandatory "class" prefix
 haftmann parents: 
36176diff
changeset | 1373 | show "class.order (le_multiset :: 'a multiset \<Rightarrow> _) less_multiset" proof | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1374 | qed (auto simp add: le_multiset_def irrefl dest: trans) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1375 | qed | 
| 10249 | 1376 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1377 | lemma mult_less_irrefl [elim!]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1378 | "M \<subset># (M::'a::order multiset) ==> R" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1379 | by (simp add: multiset_order.less_irrefl) | 
| 26567 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 1380 | |
| 10249 | 1381 | |
| 1382 | subsubsection {* Monotonicity of multiset union *}
 | |
| 1383 | ||
| 17161 | 1384 | lemma mult1_union: | 
| 40249 
cd404ecb9107
Remove unnecessary premise of mult1_union
 Lars Noschinski <noschinl@in.tum.de> parents: 
39533diff
changeset | 1385 | "(B, D) \<in> mult1 r ==> (C + B, C + D) \<in> mult1 r" | 
| 26178 | 1386 | apply (unfold mult1_def) | 
| 1387 | apply auto | |
| 1388 | apply (rule_tac x = a in exI) | |
| 1389 | apply (rule_tac x = "C + M0" in exI) | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1390 | apply (simp add: add_assoc) | 
| 26178 | 1391 | done | 
| 10249 | 1392 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1393 | lemma union_less_mono2: "B \<subset># D ==> C + B \<subset># C + (D::'a::order multiset)" | 
| 26178 | 1394 | apply (unfold less_multiset_def mult_def) | 
| 1395 | apply (erule trancl_induct) | |
| 40249 
cd404ecb9107
Remove unnecessary premise of mult1_union
 Lars Noschinski <noschinl@in.tum.de> parents: 
39533diff
changeset | 1396 | apply (blast intro: mult1_union) | 
| 
cd404ecb9107
Remove unnecessary premise of mult1_union
 Lars Noschinski <noschinl@in.tum.de> parents: 
39533diff
changeset | 1397 | apply (blast intro: mult1_union trancl_trans) | 
| 26178 | 1398 | done | 
| 10249 | 1399 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1400 | lemma union_less_mono1: "B \<subset># D ==> B + C \<subset># D + (C::'a::order multiset)" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1401 | apply (subst add_commute [of B C]) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1402 | apply (subst add_commute [of D C]) | 
| 26178 | 1403 | apply (erule union_less_mono2) | 
| 1404 | done | |
| 10249 | 1405 | |
| 17161 | 1406 | lemma union_less_mono: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1407 | "A \<subset># C ==> B \<subset># D ==> A + B \<subset># C + (D::'a::order multiset)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1408 | by (blast intro!: union_less_mono1 union_less_mono2 multiset_order.less_trans) | 
| 10249 | 1409 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1410 | interpretation multiset_order: ordered_ab_semigroup_add plus le_multiset less_multiset | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1411 | proof | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1412 | qed (auto simp add: le_multiset_def intro: union_less_mono2) | 
| 26145 | 1413 | |
| 15072 | 1414 | |
| 25610 | 1415 | subsection {* The fold combinator *}
 | 
| 1416 | ||
| 26145 | 1417 | text {*
 | 
| 1418 | The intended behaviour is | |
| 1419 |   @{text "fold_mset f z {#x\<^isub>1, ..., x\<^isub>n#} = f x\<^isub>1 (\<dots> (f x\<^isub>n z)\<dots>)"}
 | |
| 1420 |   if @{text f} is associative-commutative. 
 | |
| 25610 | 1421 | *} | 
| 1422 | ||
| 26145 | 1423 | text {*
 | 
| 1424 |   The graph of @{text "fold_mset"}, @{text "z"}: the start element,
 | |
| 1425 |   @{text "f"}: folding function, @{text "A"}: the multiset, @{text
 | |
| 1426 | "y"}: the result. | |
| 1427 | *} | |
| 25610 | 1428 | inductive | 
| 25759 | 1429 |   fold_msetG :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b \<Rightarrow> bool" 
 | 
| 25610 | 1430 | for f :: "'a \<Rightarrow> 'b \<Rightarrow> 'b" | 
| 1431 | and z :: 'b | |
| 1432 | where | |
| 25759 | 1433 |   emptyI [intro]:  "fold_msetG f z {#} z"
 | 
| 1434 | | insertI [intro]: "fold_msetG f z A y \<Longrightarrow> fold_msetG f z (A + {#x#}) (f x y)"
 | |
| 25610 | 1435 | |
| 25759 | 1436 | inductive_cases empty_fold_msetGE [elim!]: "fold_msetG f z {#} x"
 | 
| 1437 | inductive_cases insert_fold_msetGE: "fold_msetG f z (A + {#}) y" 
 | |
| 25610 | 1438 | |
| 1439 | definition | |
| 26145 | 1440 |   fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b" where
 | 
| 1441 | "fold_mset f z A = (THE x. fold_msetG f z A x)" | |
| 25610 | 1442 | |
| 25759 | 1443 | lemma Diff1_fold_msetG: | 
| 26145 | 1444 |   "fold_msetG f z (A - {#x#}) y \<Longrightarrow> x \<in># A \<Longrightarrow> fold_msetG f z A (f x y)"
 | 
| 26178 | 1445 | apply (frule_tac x = x in fold_msetG.insertI) | 
| 1446 | apply auto | |
| 1447 | done | |
| 25610 | 1448 | |
| 25759 | 1449 | lemma fold_msetG_nonempty: "\<exists>x. fold_msetG f z A x" | 
| 26178 | 1450 | apply (induct A) | 
| 1451 | apply blast | |
| 1452 | apply clarsimp | |
| 1453 | apply (drule_tac x = x in fold_msetG.insertI) | |
| 1454 | apply auto | |
| 1455 | done | |
| 25610 | 1456 | |
| 25759 | 1457 | lemma fold_mset_empty[simp]: "fold_mset f z {#} = z"
 | 
| 26178 | 1458 | unfolding fold_mset_def by blast | 
| 25610 | 1459 | |
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42809diff
changeset | 1460 | context comp_fun_commute | 
| 26145 | 1461 | begin | 
| 25610 | 1462 | |
| 26145 | 1463 | lemma fold_msetG_determ: | 
| 1464 | "fold_msetG f z A x \<Longrightarrow> fold_msetG f z A y \<Longrightarrow> y = x" | |
| 25610 | 1465 | proof (induct arbitrary: x y z rule: full_multiset_induct) | 
| 1466 | case (less M x\<^isub>1 x\<^isub>2 Z) | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1467 | have IH: "\<forall>A. A < M \<longrightarrow> | 
| 25759 | 1468 | (\<forall>x x' x''. fold_msetG f x'' A x \<longrightarrow> fold_msetG f x'' A x' | 
| 25610 | 1469 | \<longrightarrow> x' = x)" by fact | 
| 25759 | 1470 | 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 | 1471 | show ?case | 
| 25759 | 1472 | proof (rule fold_msetG.cases [OF Mfoldx\<^isub>1]) | 
| 25610 | 1473 |     assume "M = {#}" and "x\<^isub>1 = Z"
 | 
| 26145 | 1474 | then show ?case using Mfoldx\<^isub>2 by auto | 
| 25610 | 1475 | next | 
| 1476 | fix B b u | |
| 25759 | 1477 |     assume "M = B + {#b#}" and "x\<^isub>1 = f b u" and Bu: "fold_msetG f Z B u"
 | 
| 26145 | 1478 |     then have MBb: "M = B + {#b#}" and x\<^isub>1: "x\<^isub>1 = f b u" by auto
 | 
| 25610 | 1479 | show ?case | 
| 25759 | 1480 | proof (rule fold_msetG.cases [OF Mfoldx\<^isub>2]) | 
| 25610 | 1481 |       assume "M = {#}" "x\<^isub>2 = Z"
 | 
| 26145 | 1482 | then show ?case using Mfoldx\<^isub>1 by auto | 
| 25610 | 1483 | next | 
| 1484 | fix C c v | |
| 25759 | 1485 |       assume "M = C + {#c#}" and "x\<^isub>2 = f c v" and Cv: "fold_msetG f Z C v"
 | 
| 26145 | 1486 |       then have MCc: "M = C + {#c#}" and x\<^isub>2: "x\<^isub>2 = f c v" by auto
 | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1487 | then have CsubM: "C < M" by simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1488 | from MBb have BsubM: "B < M" by simp | 
| 25610 | 1489 | show ?case | 
| 1490 | proof cases | |
| 1491 | assume "b=c" | |
| 1492 | then moreover have "B = C" using MBb MCc by auto | |
| 1493 | ultimately show ?thesis using Bu Cv x\<^isub>1 x\<^isub>2 CsubM IH by auto | |
| 1494 | next | |
| 1495 | assume diff: "b \<noteq> c" | |
| 1496 |         let ?D = "B - {#c#}"
 | |
| 1497 | have cinB: "c \<in># B" and binC: "b \<in># C" using MBb MCc diff | |
| 1498 | by (auto intro: insert_noteq_member dest: sym) | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1499 |         have "B - {#c#} < B" using cinB by (rule mset_less_diff_self)
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1500 | then have DsubM: "?D < M" using BsubM by (blast intro: order_less_trans) | 
| 25610 | 1501 |         from MBb MCc have "B + {#b#} = C + {#c#}" by blast
 | 
| 26145 | 1502 |         then have [simp]: "B + {#b#} - {#c#} = C"
 | 
| 25610 | 1503 | using MBb MCc binC cinB by auto | 
| 1504 |         have B: "B = ?D + {#c#}" and C: "C = ?D + {#b#}"
 | |
| 1505 | using MBb MCc diff binC cinB | |
| 1506 | by (auto simp: multiset_add_sub_el_shuffle) | |
| 25759 | 1507 | then obtain d where Dfoldd: "fold_msetG f Z ?D d" | 
| 1508 | using fold_msetG_nonempty by iprover | |
| 26145 | 1509 | then have "fold_msetG f Z B (f c d)" using cinB | 
| 25759 | 1510 | by (rule Diff1_fold_msetG) | 
| 26145 | 1511 | then have "f c d = u" using IH BsubM Bu by blast | 
| 25610 | 1512 | moreover | 
| 25759 | 1513 | have "fold_msetG f Z C (f b d)" using binC cinB diff Dfoldd | 
| 25610 | 1514 | by (auto simp: multiset_add_sub_el_shuffle | 
| 25759 | 1515 | dest: fold_msetG.insertI [where x=b]) | 
| 26145 | 1516 | then have "f b d = v" using IH CsubM Cv by blast | 
| 25610 | 1517 | ultimately show ?thesis using x\<^isub>1 x\<^isub>2 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1518 | by (auto simp: fun_left_comm) | 
| 25610 | 1519 | qed | 
| 1520 | qed | |
| 1521 | qed | |
| 1522 | qed | |
| 1523 | ||
| 26145 | 1524 | lemma fold_mset_insert_aux: | 
| 1525 |   "(fold_msetG f z (A + {#x#}) v) =
 | |
| 25759 | 1526 | (\<exists>y. fold_msetG f z A y \<and> v = f x y)" | 
| 26178 | 1527 | apply (rule iffI) | 
| 1528 | prefer 2 | |
| 1529 | apply blast | |
| 1530 | apply (rule_tac A=A and f=f in fold_msetG_nonempty [THEN exE, standard]) | |
| 1531 | apply (blast intro: fold_msetG_determ) | |
| 1532 | done | |
| 25610 | 1533 | |
| 26145 | 1534 | lemma fold_mset_equality: "fold_msetG f z A y \<Longrightarrow> fold_mset f z A = y" | 
| 26178 | 1535 | unfolding fold_mset_def by (blast intro: fold_msetG_determ) | 
| 25610 | 1536 | |
| 26145 | 1537 | lemma fold_mset_insert: | 
| 26178 | 1538 |   "fold_mset f z (A + {#x#}) = f x (fold_mset f z A)"
 | 
| 1539 | apply (simp add: fold_mset_def fold_mset_insert_aux) | |
| 1540 | apply (rule the_equality) | |
| 1541 | apply (auto cong add: conj_cong | |
| 26145 | 1542 | simp add: fold_mset_def [symmetric] fold_mset_equality fold_msetG_nonempty) | 
| 26178 | 1543 | done | 
| 25610 | 1544 | |
| 26145 | 1545 | lemma fold_mset_commute: "f x (fold_mset f z A) = fold_mset f (f x z) A" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1546 | by (induct A) (auto simp: fold_mset_insert fun_left_comm [of x]) | 
| 26178 | 1547 | |
| 26145 | 1548 | lemma fold_mset_single [simp]: "fold_mset f z {#x#} = f x z"
 | 
| 26178 | 1549 | using fold_mset_insert [of z "{#}"] by simp
 | 
| 25610 | 1550 | |
| 26145 | 1551 | lemma fold_mset_union [simp]: | 
| 1552 | "fold_mset f z (A+B) = fold_mset f (fold_mset f z A) B" | |
| 25759 | 1553 | proof (induct A) | 
| 26145 | 1554 | case empty then show ?case by simp | 
| 25759 | 1555 | next | 
| 26145 | 1556 | case (add A x) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1557 |   have "A + {#x#} + B = (A+B) + {#x#}" by (simp add: add_ac)
 | 
| 26145 | 1558 |   then have "fold_mset f z (A + {#x#} + B) = f x (fold_mset f z (A + B))" 
 | 
| 1559 | by (simp add: fold_mset_insert) | |
| 1560 |   also have "\<dots> = fold_mset f (fold_mset f z (A + {#x#})) B"
 | |
| 1561 | by (simp add: fold_mset_commute[of x,symmetric] add fold_mset_insert) | |
| 1562 | finally show ?case . | |
| 25759 | 1563 | qed | 
| 1564 | ||
| 26145 | 1565 | lemma fold_mset_fusion: | 
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42809diff
changeset | 1566 | assumes "comp_fun_commute g" | 
| 27611 | 1567 | 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") | 
| 1568 | proof - | |
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42809diff
changeset | 1569 | interpret comp_fun_commute g by (fact assms) | 
| 27611 | 1570 | show "PROP ?P" by (induct A) auto | 
| 1571 | qed | |
| 25610 | 1572 | |
| 26145 | 1573 | lemma fold_mset_rec: | 
| 1574 | assumes "a \<in># A" | |
| 25759 | 1575 |   shows "fold_mset f z A = f a (fold_mset f z (A - {#a#}))"
 | 
| 25610 | 1576 | proof - | 
| 26145 | 1577 |   from assms obtain A' where "A = A' + {#a#}"
 | 
| 1578 | by (blast dest: multi_member_split) | |
| 1579 | then show ?thesis by simp | |
| 25610 | 1580 | qed | 
| 1581 | ||
| 26145 | 1582 | end | 
| 1583 | ||
| 1584 | text {*
 | |
| 1585 | A note on code generation: When defining some function containing a | |
| 1586 |   subterm @{term"fold_mset F"}, code generation is not automatic. When
 | |
| 1587 |   interpreting locale @{text left_commutative} with @{text F}, the
 | |
| 1588 |   would be code thms for @{const fold_mset} become thms like
 | |
| 1589 |   @{term"fold_mset F z {#} = z"} where @{text F} is not a pattern but
 | |
| 1590 | contains defined symbols, i.e.\ is not a code thm. Hence a separate | |
| 1591 |   constant with its own code thms needs to be introduced for @{text
 | |
| 1592 | F}. See the image operator below. | |
| 1593 | *} | |
| 1594 | ||
| 26016 | 1595 | |
| 1596 | subsection {* Image *}
 | |
| 1597 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1598 | definition image_mset :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1599 |   "image_mset f = fold_mset (op + o single o f) {#}"
 | 
| 26016 | 1600 | |
| 42871 
1c0b99f950d9
names of fold_set locales resemble name of characteristic property more closely
 haftmann parents: 
42809diff
changeset | 1601 | interpretation image_fun_commute: comp_fun_commute "op + o single o f" | 
| 42809 
5b45125b15ba
use pointfree characterisation for fold_set locale
 haftmann parents: 
41505diff
changeset | 1602 | proof qed (simp add: add_ac fun_eq_iff) | 
| 26016 | 1603 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1604 | lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
 | 
| 26178 | 1605 | by (simp add: image_mset_def) | 
| 26016 | 1606 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1607 | lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}"
 | 
| 26178 | 1608 | by (simp add: image_mset_def) | 
| 26016 | 1609 | |
| 1610 | lemma image_mset_insert: | |
| 1611 |   "image_mset f (M + {#a#}) = image_mset f M + {#f a#}"
 | |
| 26178 | 1612 | by (simp add: image_mset_def add_ac) | 
| 26016 | 1613 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1614 | lemma image_mset_union [simp]: | 
| 26016 | 1615 | "image_mset f (M+N) = image_mset f M + image_mset f N" | 
| 26178 | 1616 | apply (induct N) | 
| 1617 | apply simp | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1618 | apply (simp add: add_assoc [symmetric] image_mset_insert) | 
| 26178 | 1619 | done | 
| 26016 | 1620 | |
| 26145 | 1621 | lemma size_image_mset [simp]: "size (image_mset f M) = size M" | 
| 26178 | 1622 | by (induct M) simp_all | 
| 26016 | 1623 | |
| 26145 | 1624 | lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \<longleftrightarrow> M = {#}"
 | 
| 26178 | 1625 | by (cases M) auto | 
| 26016 | 1626 | |
| 26145 | 1627 | syntax | 
| 35352 | 1628 | "_comprehension1_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" | 
| 26145 | 1629 |       ("({#_/. _ :# _#})")
 | 
| 1630 | translations | |
| 1631 |   "{#e. x:#M#}" == "CONST image_mset (%x. e) M"
 | |
| 26016 | 1632 | |
| 26145 | 1633 | syntax | 
| 35352 | 1634 | "_comprehension2_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" | 
| 26145 | 1635 |       ("({#_/ | _ :# _./ _#})")
 | 
| 26016 | 1636 | translations | 
| 26033 | 1637 |   "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}"
 | 
| 26016 | 1638 | |
| 26145 | 1639 | text {*
 | 
| 1640 |   This allows to write not just filters like @{term "{#x:#M. x<c#}"}
 | |
| 1641 |   but also images like @{term "{#x+x. x:#M #}"} and @{term [source]
 | |
| 1642 |   "{#x+x|x:#M. x<c#}"}, where the latter is currently displayed as
 | |
| 1643 |   @{term "{#x+x|x:#M. x<c#}"}.
 | |
| 1644 | *} | |
| 26016 | 1645 | |
| 41505 
6d19301074cf
"enriched_type" replaces less specific "type_lifting"
 haftmann parents: 
41372diff
changeset | 1646 | enriched_type image_mset: image_mset proof - | 
| 41372 | 1647 | fix f g | 
| 1648 | show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)" | |
| 1649 | proof | |
| 1650 | fix A | |
| 1651 | show "(image_mset f \<circ> image_mset g) A = image_mset (f \<circ> g) A" | |
| 1652 | by (induct A) simp_all | |
| 1653 | qed | |
| 40606 | 1654 | next | 
| 41372 | 1655 | show "image_mset id = id" | 
| 1656 | proof | |
| 1657 | fix A | |
| 1658 | show "image_mset id A = id A" | |
| 1659 | by (induct A) simp_all | |
| 1660 | qed | |
| 40606 | 1661 | qed | 
| 1662 | ||
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1663 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1664 | subsection {* Termination proofs with multiset orders *}
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1665 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1666 | 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 | 1667 |   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 | 1668 |   and multi_member_last: "x \<in># {# x #}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1669 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1670 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1671 | definition "ms_strict = mult pair_less" | 
| 37765 | 1672 | definition "ms_weak = ms_strict \<union> Id" | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1673 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1674 | 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 | 1675 | 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 | 1676 | 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 | 1677 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1678 | lemma smsI: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1679 | "(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 | 1680 | unfolding ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1681 | 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 | 1682 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1683 | lemma wmsI: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1684 |   "(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 | 1685 | \<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 | 1686 | unfolding ms_weak_def ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1687 | 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 | 1688 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1689 | inductive pw_leq | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1690 | where | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1691 |   pw_leq_empty: "pw_leq {#} {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1692 | | 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 | 1693 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1694 | lemma pw_leq_lstep: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1695 |   "(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 | 1696 | 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 | 1697 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1698 | lemma pw_leq_split: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1699 | assumes "pw_leq X Y" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1700 |   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 | 1701 | using assms | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1702 | proof (induct) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1703 | 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 | 1704 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1705 | 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 | 1706 | then obtain A B Z where | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1707 | [simp]: "X = A + Z" "Y = B + Z" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1708 |       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 | 1709 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1710 | 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 | 1711 | unfolding pair_leq_def by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1712 | thus ?case | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1713 | proof | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1714 | assume [simp]: "x = y" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1715 | have | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1716 |       "{#x#} + X = A + ({#y#}+Z) 
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1717 |       \<and> {#y#} + Y = B + ({#y#}+Z)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1718 |       \<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 | 1719 | by (auto simp: add_ac) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1720 | thus ?case by (intro exI) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1721 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1722 | assume A: "(x, y) \<in> pair_less" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1723 |     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 | 1724 |     have "{#x#} + X = ?A' + Z"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1725 |       "{#y#} + Y = ?B' + Z"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1726 | by (auto simp add: add_ac) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1727 | moreover have | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1728 | "(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 | 1729 | using 1 A unfolding max_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1730 | by (auto elim!: max_ext.cases) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1731 | ultimately show ?thesis by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1732 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1733 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1734 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1735 | lemma | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1736 | assumes pwleq: "pw_leq Z Z'" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1737 | 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 | 1738 | 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 | 1739 |   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 | 1740 | proof - | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1741 | from pw_leq_split[OF pwleq] | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1742 | obtain A' B' Z'' | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1743 | 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 | 1744 |     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 | 1745 | by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1746 |   {
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1747 | 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 | 1748 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1749 | 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 | 1750 | proof | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1751 | 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 | 1752 | 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 | 1753 | 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 | 1754 | thus ?thesis by (rule smsI) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1755 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1756 |       assume [simp]: "A' = {#} \<and> B' = {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1757 | 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 | 1758 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1759 | 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 | 1760 | 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 | 1761 | } | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1762 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1763 | 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 | 1764 |   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 | 1765 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1766 | |
| 39301 | 1767 | lemma empty_neutral: "{#} + x = x" "x + {#} = x"
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1768 | and nonempty_plus: "{# x #} + rs \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1769 | and nonempty_single: "{# x #} \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1770 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1771 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1772 | setup {*
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1773 | let | 
| 35402 | 1774 |   fun msetT T = Type (@{type_name multiset}, [T]);
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1775 | |
| 35402 | 1776 |   fun mk_mset T [] = Const (@{const_abbrev Mempty}, msetT T)
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1777 |     | 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 | 1778 | | mk_mset T (x :: xs) = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1779 |           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 | 1780 | 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 | 1781 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1782 | fun mset_member_tac m i = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1783 | (if m <= 0 then | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1784 |            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 | 1785 | else | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1786 |            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 | 1787 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1788 | val mset_nonempty_tac = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1789 |       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 | 1790 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1791 | val regroup_munion_conv = | 
| 35402 | 1792 |       Function_Lib.regroup_conv @{const_abbrev Mempty} @{const_name plus}
 | 
| 39301 | 1793 |         (map (fn t => t RS eq_reflection) (@{thms add_ac} @ @{thms empty_neutral}))
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1794 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1795 | fun unfold_pwleq_tac i = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1796 |     (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 | 1797 |       ORELSE (rtac @{thm pw_leq_lstep} i)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1798 |       ORELSE (rtac @{thm pw_leq_empty} i)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1799 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1800 |   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 | 1801 |                       @{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 | 1802 | in | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1803 | ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1804 |   {
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1805 | 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 | 1806 | 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 | 1807 | 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 | 1808 |     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 | 1809 |     reduction_pair= @{thm ms_reduction_pair}
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1810 | }) | 
| 10249 | 1811 | end | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1812 | *} | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1813 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1814 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1815 | subsection {* Legacy theorem bindings *}
 | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1816 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1817 | lemmas multi_count_eq = multiset_eq_iff [symmetric] | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1818 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1819 | lemma union_commute: "M + N = N + (M::'a multiset)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1820 | by (fact add_commute) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1821 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1822 | lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1823 | by (fact add_assoc) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1824 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1825 | lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1826 | by (fact add_left_commute) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1827 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1828 | lemmas union_ac = union_assoc union_commute union_lcomm | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1829 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1830 | lemma union_right_cancel: "M + K = N + K \<longleftrightarrow> M = (N::'a multiset)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1831 | by (fact add_right_cancel) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1832 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1833 | lemma union_left_cancel: "K + M = K + N \<longleftrightarrow> M = (N::'a multiset)" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1834 | by (fact add_left_cancel) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1835 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1836 | lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y" | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1837 | by (fact add_imp_eq) | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1838 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1839 | lemma mset_less_trans: "(M::'a multiset) < K \<Longrightarrow> K < N \<Longrightarrow> M < N" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1840 | by (fact order_less_trans) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1841 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1842 | lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1843 | by (fact inf.commute) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1844 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1845 | lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1846 | by (fact inf.assoc [symmetric]) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1847 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1848 | lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1849 | by (fact inf.left_commute) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1850 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1851 | lemmas multiset_inter_ac = | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1852 | multiset_inter_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1853 | multiset_inter_assoc | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1854 | multiset_inter_left_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1855 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1856 | lemma mult_less_not_refl: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1857 | "\<not> M \<subset># (M::'a::order multiset)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1858 | by (fact multiset_order.less_irrefl) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1859 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1860 | lemma mult_less_trans: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1861 | "K \<subset># M ==> M \<subset># N ==> K \<subset># (N::'a::order multiset)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1862 | by (fact multiset_order.less_trans) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1863 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1864 | lemma mult_less_not_sym: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1865 | "M \<subset># N ==> \<not> N \<subset># (M::'a::order multiset)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1866 | by (fact multiset_order.less_not_sym) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1867 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1868 | lemma mult_less_asym: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1869 | "M \<subset># N ==> (\<not> P ==> N \<subset># (M::'a::order multiset)) ==> P" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1870 | by (fact multiset_order.less_asym) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1871 | |
| 35712 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1872 | ML {*
 | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1873 | fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1874 | (Const _ $ t') = | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1875 | let | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1876 | val (maybe_opt, ps) = | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1877 | Nitpick_Model.dest_plain_fun t' ||> op ~~ | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1878 | ||> map (apsnd (snd o HOLogic.dest_number)) | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1879 | fun elems_for t = | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1880 | case AList.lookup (op =) ps t of | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1881 | SOME n => replicate n t | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1882 | | NONE => [Const (maybe_name, elem_T --> elem_T) $ t] | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1883 | in | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1884 | case maps elems_for (all_values elem_T) @ | 
| 37261 | 1885 | (if maybe_opt then [Const (Nitpick_Model.unrep (), elem_T)] | 
| 1886 | else []) of | |
| 35712 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1887 |         [] => Const (@{const_name zero_class.zero}, T)
 | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1888 | | ts => foldl1 (fn (t1, t2) => | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1889 |                          Const (@{const_name plus_class.plus}, T --> T --> T)
 | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1890 | $ t1 $ t2) | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1891 |                      (map (curry (op $) (Const (@{const_name single},
 | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1892 | elem_T --> T))) ts) | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1893 | end | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1894 | | multiset_postproc _ _ _ _ t = t | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1895 | *} | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1896 | |
| 38287 | 1897 | declaration {*
 | 
| 1898 | Nitpick_Model.register_term_postprocessor @{typ "'a multiset"}
 | |
| 38242 | 1899 | multiset_postproc | 
| 35712 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1900 | *} | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 1901 | |
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
37107diff
changeset | 1902 | end |