| author | paulson <lp15@cam.ac.uk> | 
| Fri, 20 Mar 2015 16:11:28 +0000 | |
| changeset 59765 | 26d1c71784f1 | 
| parent 59625 | aacdce52b2fc | 
| child 59807 | 22bc39064290 | 
| permissions | -rw-r--r-- | 
| 10249 | 1 | (* Title: HOL/Library/Multiset.thy | 
| 15072 | 2 | Author: Tobias Nipkow, Markus Wenzel, Lawrence C Paulson, Norbert Voelker | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 3 | Author: Andrei Popescu, TU Muenchen | 
| 10249 | 4 | *) | 
| 5 | ||
| 58881 | 6 | section {* (Finite) multisets *}
 | 
| 10249 | 7 | |
| 15131 | 8 | theory Multiset | 
| 51599 
1559e9266280
optionalized very specific code setup for multisets
 haftmann parents: 
51548diff
changeset | 9 | imports Main | 
| 15131 | 10 | begin | 
| 10249 | 11 | |
| 12 | subsection {* The type of multisets *}
 | |
| 13 | ||
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45608diff
changeset | 14 | definition "multiset = {f :: 'a => nat. finite {x. f x > 0}}"
 | 
| 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45608diff
changeset | 15 | |
| 49834 | 16 | typedef 'a multiset = "multiset :: ('a => nat) 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 | 17 | morphisms count Abs_multiset | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45608diff
changeset | 18 | unfolding multiset_def | 
| 10249 | 19 | proof | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45608diff
changeset | 20 |   show "(\<lambda>x. 0::nat) \<in> {f. finite {x. f x > 0}}" by simp
 | 
| 10249 | 21 | qed | 
| 22 | ||
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 23 | setup_lifting type_definition_multiset | 
| 19086 | 24 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 25 | abbreviation Melem :: "'a => 'a multiset => bool"  ("(_/ :# _)" [50, 51] 50) where
 | 
| 25610 | 26 | "a :# M == 0 < count M a" | 
| 27 | ||
| 26145 | 28 | notation (xsymbols) | 
| 29 | Melem (infix "\<in>#" 50) | |
| 10249 | 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_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 | 32 | "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 | 33 | 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 | 34 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 35 | 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 | 36 | "(\<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 | 37 | 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 | 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 | 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 | 40 |  \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 | 41 | *} | 
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 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 | 44 | "(\<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 | 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 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 | 48 | "(\<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 | 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 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 | 52 | "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 | 53 | 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 | 54 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 56 | 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 | 57 | 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 | 58 | 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 | 59 |   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 | 60 | 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 | 61 | 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 | 62 | 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 | 63 | 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 | 64 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 65 | 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 | 66 | 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 | 67 | 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 | 68 | 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 | 69 |   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 | 70 | 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 | 71 | 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 | 72 | 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 | 73 | 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 | 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 | 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 | 76 | 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 | 77 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 80 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 81 | 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 | 82 | |
| 48008 | 83 | instantiation multiset :: (type) cancel_comm_monoid_add | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 84 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 85 | |
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 86 | lift_definition zero_multiset :: "'a multiset" is "\<lambda>a. 0" | 
| 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 87 | by (rule const0_in_multiset) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 88 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 90 | "Mempty \<equiv> 0" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 91 | |
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 92 | lift_definition plus_multiset :: "'a multiset => 'a multiset => 'a multiset" is "\<lambda>M N. (\<lambda>a. M a + N a)" | 
| 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 93 | by (rule union_preserves_multiset) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 94 | |
| 48008 | 95 | instance | 
| 96 | by default (transfer, simp add: fun_eq_iff)+ | |
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 97 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 98 | end | 
| 10249 | 99 | |
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 100 | lift_definition single :: "'a => 'a multiset" is "\<lambda>a b. if b = a then 1 else 0" | 
| 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 101 | by (rule only1_in_multiset) | 
| 15869 | 102 | |
| 26145 | 103 | syntax | 
| 26176 | 104 |   "_multiset" :: "args => 'a multiset"    ("{#(_)#}")
 | 
| 25507 | 105 | translations | 
| 106 |   "{#x, xs#}" == "{#x#} + {#xs#}"
 | |
| 107 |   "{#x#}" == "CONST single x"
 | |
| 108 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 109 | lemma count_empty [simp]: "count {#} a = 0"
 | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 110 | by (simp add: zero_multiset.rep_eq) | 
| 10249 | 111 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 112 | lemma count_single [simp]: "count {#b#} a = (if b = a then 1 else 0)"
 | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 113 | by (simp add: single.rep_eq) | 
| 29901 | 114 | |
| 10249 | 115 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | subsection {* Basic operations *}
 | 
| 10249 | 117 | |
| 118 | subsubsection {* Union *}
 | |
| 119 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 120 | lemma count_union [simp]: "count (M + N) a = count M a + count N a" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 121 | by (simp add: plus_multiset.rep_eq) | 
| 10249 | 122 | |
| 123 | ||
| 124 | subsubsection {* Difference *}
 | |
| 125 | ||
| 49388 | 126 | instantiation multiset :: (type) comm_monoid_diff | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 127 | 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 | 128 | |
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 129 | lift_definition minus_multiset :: "'a multiset => 'a multiset => 'a multiset" is "\<lambda> M N. \<lambda>a. M a - N a" | 
| 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 130 | by (rule diff_preserves_multiset) | 
| 58425 | 131 | |
| 49388 | 132 | instance | 
| 133 | by default (transfer, simp add: fun_eq_iff)+ | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 134 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 136 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 137 | lemma count_diff [simp]: "count (M - N) a = count M a - count N a" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 138 | by (simp add: minus_multiset.rep_eq) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 139 | |
| 17161 | 140 | lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
 | 
| 52289 | 141 | by rule (fact Groups.diff_zero, fact Groups.zero_diff) | 
| 36903 | 142 | |
| 143 | lemma diff_cancel[simp]: "A - A = {#}"
 | |
| 52289 | 144 | by (fact Groups.diff_cancel) | 
| 10249 | 145 | |
| 36903 | 146 | lemma diff_union_cancelR [simp]: "M + N - N = (M::'a multiset)" | 
| 52289 | 147 | by (fact add_diff_cancel_right') | 
| 10249 | 148 | |
| 36903 | 149 | lemma diff_union_cancelL [simp]: "N + M - N = (M::'a multiset)" | 
| 52289 | 150 | by (fact add_diff_cancel_left') | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 151 | |
| 52289 | 152 | lemma diff_right_commute: | 
| 153 | "(M::'a multiset) - N - Q = M - Q - N" | |
| 154 | by (fact diff_right_commute) | |
| 155 | ||
| 156 | lemma diff_add: | |
| 157 | "(M::'a multiset) - (N + Q) = M - N - Q" | |
| 158 | by (rule sym) (fact diff_diff_add) | |
| 58425 | 159 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 160 | 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 | 161 |   "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 | 162 | 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 | 163 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 165 |   "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 | 166 | 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 | 167 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 169 |   "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 | 170 | 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 | 171 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 172 | 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 | 173 |   "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 | 174 | by (simp add: multiset_eq_iff) | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 175 | |
| 10249 | 176 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 177 | 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 | 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_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 | 180 | 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 | 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 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 | 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 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 | 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 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 | 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 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 | 192 | 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 | 193 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 195 |   "\<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 | 196 | 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 | 197 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 199 |   "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 | 200 | 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 | 201 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 203 |   "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 | 204 | 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 | 205 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 207 |   "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 | 208 | 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 | 209 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma union_is_single: | 
| 46730 | 211 |   "M + N = {#a#} \<longleftrightarrow> M = {#a#} \<and> N={#} \<or> M = {#} \<and> N = {#a#}" (is "?lhs = ?rhs")
 | 
| 212 | 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 | 213 | 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 | 214 | next | 
| 46730 | 215 | assume ?lhs then show ?rhs | 
| 216 | 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 | 217 | 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 | 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 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 | 220 |   "{#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 | 221 |   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 | 222 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 224 |   "M + {#a#} = N + {#b#} \<longleftrightarrow> M = N \<and> a = b \<or> M = N - {#a#} + {#b#} \<and> N = M - {#b#} + {#a#}"  (is "?lhs = ?rhs")
 | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44339diff
changeset | 225 | (* shorter: by (simp add: multiset_eq_iff) fastforce *) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 226 | 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 | 227 | assume ?rhs then show ?lhs | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 228 |   by (auto simp add: add.assoc add.commute [of "{#b#}"])
 | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 229 | (drule sym, simp add: add.assoc [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 | 230 | 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 | 231 | 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 | 232 | 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 | 233 | 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 | 234 | 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 | 235 | 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 | 236 | 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 | 237 |     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 | 238 | 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 | 239 |     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 | 240 | 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 | 241 |     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 | 242 | 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 | 243 | 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 | 244 | |
| 58425 | 245 | lemma insert_noteq_member: | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 246 |   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 | 247 | 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 | 248 | 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 | 249 | 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 | 250 |   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 | 251 |   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 | 252 |   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 | 253 | 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 | 254 | 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 | 255 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 257 |   "(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 | 258 |     (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 | 259 | 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 | 260 | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 261 | lemma multi_member_split: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 262 |   "x \<in># M \<Longrightarrow> \<exists>A. M = A + {#x#}"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 263 |   by (rule_tac x = "M - {#x#}" in exI, simp)
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 264 | |
| 58425 | 265 | lemma multiset_add_sub_el_shuffle: | 
| 266 | assumes "c \<in># B" and "b \<noteq> c" | |
| 58098 | 267 |   shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}"
 | 
| 268 | proof - | |
| 58425 | 269 |   from `c \<in># B` obtain A where B: "B = A + {#c#}"
 | 
| 58098 | 270 | by (blast dest: multi_member_split) | 
| 271 |   have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp
 | |
| 58425 | 272 |   then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}"
 | 
| 58098 | 273 | by (simp add: ac_simps) | 
| 274 | then show ?thesis using B by simp | |
| 275 | qed | |
| 276 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 278 | 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 | 279 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 280 | instantiation multiset :: (type) ordered_ab_semigroup_add_imp_le | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 281 | begin | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 282 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55467diff
changeset | 283 | lift_definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" is "\<lambda> A B. (\<forall>a. A a \<le> B a)" . | 
| 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55467diff
changeset | 284 | |
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 285 | lemmas mset_le_def = less_eq_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 | 286 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 287 | definition less_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 288 | 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 | 289 | |
| 46921 | 290 | instance | 
| 291 | by default (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 | 292 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 293 | 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 | 294 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma mset_less_eqI: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 296 | "(\<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 | 297 | 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 | 298 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 299 | lemma mset_le_exists_conv: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 300 | "(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 | 301 | 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 | 302 | 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 | 303 | 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 | 304 | |
| 52289 | 305 | instance multiset :: (type) ordered_cancel_comm_monoid_diff | 
| 306 | by default (simp, fact mset_le_exists_conv) | |
| 307 | ||
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 308 | lemma mset_le_mono_add_right_cancel [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 309 | "(A::'a multiset) + C \<le> B + C \<longleftrightarrow> A \<le> B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 310 | 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 | 311 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 312 | lemma mset_le_mono_add_left_cancel [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 313 | "C + (A::'a multiset) \<le> C + B \<longleftrightarrow> A \<le> B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 314 | by (fact add_le_cancel_left) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 315 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 316 | lemma mset_le_mono_add: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 317 | "(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 | 318 | 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 | 319 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 320 | lemma mset_le_add_left [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 321 | "(A::'a multiset) \<le> A + B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 322 | unfolding mset_le_def by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 323 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 324 | lemma mset_le_add_right [simp]: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 325 | "B \<le> (A::'a multiset) + B" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 326 | 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 | 327 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 328 | lemma mset_le_single: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 329 |   "a :# B \<Longrightarrow> {#a#} \<le> B"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 330 | 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 | 331 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 332 | lemma multiset_diff_union_assoc: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 333 | "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 | 334 | 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 | 335 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma mset_le_multiset_union_diff_commute: | 
| 36867 | 337 | "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 | 338 | 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 | 339 | |
| 39301 | 340 | lemma diff_le_self[simp]: "(M::'a multiset) - N \<le> M" | 
| 341 | by(simp add: mset_le_def) | |
| 342 | ||
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 343 | 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 | 344 | 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 | 345 | 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 | 346 | 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 | 347 | 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 | 348 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 349 | 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 | 350 | 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 | 351 | 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 | 352 | 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 | 353 | done | 
| 58425 | 354 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 355 | 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 | 356 | 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 | 357 | 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 | 358 | 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 | 359 | 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 | 360 | 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 | 361 | 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 | 362 | 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 | 363 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 364 | 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 | 365 | 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 | 366 | 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 | 367 | 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 | 368 | 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 | 369 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 370 | 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 | 371 | 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 | 372 | |
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 373 | lemma empty_le[simp]: "{#} \<le> A"
 | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 374 | unfolding mset_le_exists_conv by auto | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 375 | |
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 376 | lemma le_empty[simp]: "(M \<le> {#}) = (M = {#})"
 | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 377 | unfolding mset_le_exists_conv by auto | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 378 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 379 | lemma multi_psub_of_add_self[simp]: "A < A + {#x#}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 380 | 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 | 381 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 382 | lemma multi_psub_self[simp]: "(A::'a multiset) < A = False" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 383 | 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 | 384 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 385 | lemma mset_less_add_bothsides: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 386 |   "T + {#x#} < S + {#x#} \<Longrightarrow> T < S"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 387 | by (fact add_less_imp_less_right) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 388 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 389 | lemma mset_less_empty_nonempty: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 390 |   "{#} < S \<longleftrightarrow> S \<noteq> {#}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 391 | by (auto simp: mset_le_def mset_less_def) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 392 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 393 | lemma mset_less_diff_self: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 394 |   "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 | 395 | 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 | 396 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 397 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 398 | subsubsection {* Intersection *}
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 399 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 400 | instantiation multiset :: (type) semilattice_inf | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 401 | begin | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 402 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 403 | definition inf_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 404 | multiset_inter_def: "inf_multiset A B = A - (A - B)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 405 | |
| 46921 | 406 | instance | 
| 407 | proof - | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 408 | have aux: "\<And>m n q :: nat. m \<le> n \<Longrightarrow> m \<le> q \<Longrightarrow> m \<le> n - (n - q)" by arith | 
| 46921 | 409 |   show "OFCLASS('a multiset, semilattice_inf_class)"
 | 
| 410 | by default (auto simp add: multiset_inter_def mset_le_def aux) | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 411 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 412 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 413 | end | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 414 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 415 | 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 | 416 | "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 | 417 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 418 | lemma multiset_inter_count [simp]: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 419 | "count (A #\<inter> B) x = min (count A x) (count B x)" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 420 | by (simp add: multiset_inter_def) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 421 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 422 | lemma multiset_inter_single: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
 | 
| 46730 | 423 | by (rule multiset_eqI) auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 424 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 425 | lemma multiset_union_diff_commute: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 426 |   assumes "B #\<inter> C = {#}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 427 | 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 | 428 | proof (rule multiset_eqI) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 429 | fix x | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 430 | from assms have "min (count B x) (count C x) = 0" | 
| 46730 | 431 | by (auto simp add: multiset_eq_iff) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 432 | then have "count B x = 0 \<or> count C x = 0" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 433 | by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 434 | then show "count (A + B - C) x = count (A - C + B) x" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 435 | by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 436 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 437 | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 438 | lemma empty_inter [simp]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 439 |   "{#} #\<inter> M = {#}"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 440 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 441 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 442 | lemma inter_empty [simp]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 443 |   "M #\<inter> {#} = {#}"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 444 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 445 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 446 | lemma inter_add_left1: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 447 |   "\<not> x \<in># N \<Longrightarrow> (M + {#x#}) #\<inter> N = M #\<inter> N"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 448 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 449 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 450 | lemma inter_add_left2: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 451 |   "x \<in># N \<Longrightarrow> (M + {#x#}) #\<inter> N = (M #\<inter> (N - {#x#})) + {#x#}"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 452 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 453 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 454 | lemma inter_add_right1: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 455 |   "\<not> x \<in># N \<Longrightarrow> N #\<inter> (M + {#x#}) = N #\<inter> M"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 456 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 457 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 458 | lemma inter_add_right2: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 459 |   "x \<in># N \<Longrightarrow> N #\<inter> (M + {#x#}) = ((N - {#x#}) #\<inter> M) + {#x#}"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 460 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 461 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 462 | |
| 51623 | 463 | subsubsection {* Bounded union *}
 | 
| 464 | ||
| 465 | instantiation multiset :: (type) semilattice_sup | |
| 466 | begin | |
| 467 | ||
| 468 | definition sup_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" where | |
| 469 | "sup_multiset A B = A + (B - A)" | |
| 470 | ||
| 471 | instance | |
| 472 | proof - | |
| 473 | have aux: "\<And>m n q :: nat. m \<le> n \<Longrightarrow> q \<le> n \<Longrightarrow> m + (q - m) \<le> n" by arith | |
| 474 |   show "OFCLASS('a multiset, semilattice_sup_class)"
 | |
| 475 | by default (auto simp add: sup_multiset_def mset_le_def aux) | |
| 476 | qed | |
| 477 | ||
| 478 | end | |
| 479 | ||
| 480 | abbreviation sup_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<union>" 70) where | |
| 481 | "sup_multiset \<equiv> sup" | |
| 482 | ||
| 483 | lemma sup_multiset_count [simp]: | |
| 484 | "count (A #\<union> B) x = max (count A x) (count B x)" | |
| 485 | by (simp add: sup_multiset_def) | |
| 486 | ||
| 487 | lemma empty_sup [simp]: | |
| 488 |   "{#} #\<union> M = M"
 | |
| 489 | by (simp add: multiset_eq_iff) | |
| 490 | ||
| 491 | lemma sup_empty [simp]: | |
| 492 |   "M #\<union> {#} = M"
 | |
| 493 | by (simp add: multiset_eq_iff) | |
| 494 | ||
| 495 | lemma sup_add_left1: | |
| 496 |   "\<not> x \<in># N \<Longrightarrow> (M + {#x#}) #\<union> N = (M #\<union> N) + {#x#}"
 | |
| 497 | by (simp add: multiset_eq_iff) | |
| 498 | ||
| 499 | lemma sup_add_left2: | |
| 500 |   "x \<in># N \<Longrightarrow> (M + {#x#}) #\<union> N = (M #\<union> (N - {#x#})) + {#x#}"
 | |
| 501 | by (simp add: multiset_eq_iff) | |
| 502 | ||
| 503 | lemma sup_add_right1: | |
| 504 |   "\<not> x \<in># N \<Longrightarrow> N #\<union> (M + {#x#}) = (N #\<union> M) + {#x#}"
 | |
| 505 | by (simp add: multiset_eq_iff) | |
| 506 | ||
| 507 | lemma sup_add_right2: | |
| 508 |   "x \<in># N \<Longrightarrow> N #\<union> (M + {#x#}) = ((N - {#x#}) #\<union> M) + {#x#}"
 | |
| 509 | by (simp add: multiset_eq_iff) | |
| 510 | ||
| 511 | ||
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 512 | subsubsection {* Filter (with comprehension syntax) *}
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 513 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 514 | text {* Multiset comprehension *}
 | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 515 | |
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 516 | lift_definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" is "\<lambda>P M. \<lambda>x. if P x then M x else 0"
 | 
| 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 517 | by (rule filter_preserves_multiset) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 518 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 519 | hide_const (open) filter | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 520 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 521 | lemma count_filter [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 522 | "count (Multiset.filter P M) a = (if P a then count M a else 0)" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 523 | by (simp add: filter.rep_eq) | 
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 524 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 525 | lemma filter_empty [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 526 |   "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 | 527 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 528 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 529 | lemma filter_single [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 530 |   "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 | 531 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 532 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 533 | lemma filter_union [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 534 | "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 | 535 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 536 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 537 | lemma filter_diff [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 538 | "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 | 539 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 540 | |
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 541 | lemma filter_inter [simp]: | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 542 | "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 | 543 | by (rule multiset_eqI) simp | 
| 10249 | 544 | |
| 58035 | 545 | lemma multiset_filter_subset[simp]: "Multiset.filter f M \<le> M" | 
| 546 | unfolding less_eq_multiset.rep_eq by auto | |
| 547 | ||
| 548 | lemma multiset_filter_mono: assumes "A \<le> B" | |
| 549 | shows "Multiset.filter f A \<le> Multiset.filter f B" | |
| 550 | proof - | |
| 551 | from assms[unfolded mset_le_exists_conv] | |
| 552 | obtain C where B: "B = A + C" by auto | |
| 553 | show ?thesis unfolding B by auto | |
| 554 | qed | |
| 555 | ||
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 556 | syntax | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 557 |   "_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 | 558 | syntax (xsymbol) | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 559 |   "_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 | 560 | translations | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 561 |   "{#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 | 562 | |
| 10249 | 563 | |
| 564 | subsubsection {* Set of elements *}
 | |
| 565 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 566 | 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 | 567 |   "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 | 568 | |
| 17161 | 569 | lemma set_of_empty [simp]: "set_of {#} = {}"
 | 
| 26178 | 570 | by (simp add: set_of_def) | 
| 10249 | 571 | |
| 17161 | 572 | lemma set_of_single [simp]: "set_of {#b#} = {b}"
 | 
| 26178 | 573 | by (simp add: set_of_def) | 
| 10249 | 574 | |
| 17161 | 575 | lemma set_of_union [simp]: "set_of (M + N) = set_of M \<union> set_of N" | 
| 26178 | 576 | by (auto simp add: set_of_def) | 
| 10249 | 577 | |
| 17161 | 578 | 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 | 579 | by (auto simp add: set_of_def multiset_eq_iff) | 
| 10249 | 580 | |
| 17161 | 581 | lemma mem_set_of_iff [simp]: "(x \<in> set_of M) = (x :# M)" | 
| 26178 | 582 | by (auto simp add: set_of_def) | 
| 26016 | 583 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 584 | lemma set_of_filter [simp]: "set_of {# x:#M. P x #} = set_of M \<inter> {x. P x}"
 | 
| 26178 | 585 | by (auto simp add: set_of_def) | 
| 10249 | 586 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 587 | 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 | 588 | 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 | 589 | |
| 46756 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 bulwahn parents: 
46730diff
changeset | 590 | lemma finite_Collect_mem [iff]: "finite {x. x :# M}"
 | 
| 
faf62905cd53
adding finiteness of intervals on integer sets; adding another finiteness theorem for multisets
 bulwahn parents: 
46730diff
changeset | 591 | unfolding set_of_def[symmetric] by simp | 
| 10249 | 592 | |
| 58425 | 593 | lemma set_of_mono: "A \<le> B \<Longrightarrow> set_of A \<subseteq> set_of B" | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 594 | by (metis mset_leD subsetI mem_set_of_iff) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 595 | |
| 10249 | 596 | subsubsection {* Size *}
 | 
| 597 | ||
| 56656 | 598 | definition wcount where "wcount f M = (\<lambda>x. count M x * Suc (f x))" | 
| 599 | ||
| 600 | lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" | |
| 601 | by (auto simp: wcount_def add_mult_distrib) | |
| 602 | ||
| 603 | definition size_multiset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a multiset \<Rightarrow> nat" where
 | |
| 604 | "size_multiset f M = setsum (wcount f M) (set_of M)" | |
| 605 | ||
| 606 | lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] | |
| 607 | ||
| 608 | instantiation multiset :: (type) size begin | |
| 609 | definition size_multiset where | |
| 610 | size_multiset_overloaded_def: "size_multiset = Multiset.size_multiset (\<lambda>_. 0)" | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 611 | 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 | 612 | 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 | 613 | |
| 56656 | 614 | lemmas size_multiset_overloaded_eq = | 
| 615 | size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] | |
| 616 | ||
| 617 | lemma size_multiset_empty [simp]: "size_multiset f {#} = 0"
 | |
| 618 | by (simp add: size_multiset_def) | |
| 619 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 620 | lemma size_empty [simp]: "size {#} = 0"
 | 
| 56656 | 621 | by (simp add: size_multiset_overloaded_def) | 
| 622 | ||
| 623 | lemma size_multiset_single [simp]: "size_multiset f {#b#} = Suc (f b)"
 | |
| 624 | by (simp add: size_multiset_eq) | |
| 10249 | 625 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 626 | lemma size_single [simp]: "size {#b#} = 1"
 | 
| 56656 | 627 | by (simp add: size_multiset_overloaded_def) | 
| 628 | ||
| 629 | lemma setsum_wcount_Int: | |
| 630 | "finite A \<Longrightarrow> setsum (wcount f N) (A \<inter> set_of N) = setsum (wcount f N) A" | |
| 26178 | 631 | apply (induct rule: finite_induct) | 
| 632 | apply simp | |
| 56656 | 633 | apply (simp add: Int_insert_left set_of_def wcount_def) | 
| 634 | done | |
| 635 | ||
| 636 | lemma size_multiset_union [simp]: | |
| 637 | "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" | |
| 57418 | 638 | apply (simp add: size_multiset_def setsum_Un_nat setsum.distrib setsum_wcount_Int wcount_union) | 
| 56656 | 639 | apply (subst Int_commute) | 
| 640 | apply (simp add: setsum_wcount_Int) | |
| 26178 | 641 | done | 
| 10249 | 642 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 643 | lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" | 
| 56656 | 644 | by (auto simp add: size_multiset_overloaded_def) | 
| 645 | ||
| 646 | lemma size_multiset_eq_0_iff_empty [iff]: "(size_multiset f M = 0) = (M = {#})"
 | |
| 647 | by (auto simp add: size_multiset_eq multiset_eq_iff) | |
| 10249 | 648 | |
| 17161 | 649 | lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
 | 
| 56656 | 650 | by (auto simp add: size_multiset_overloaded_def) | 
| 26016 | 651 | |
| 652 | lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
 | |
| 26178 | 653 | by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) | 
| 10249 | 654 | |
| 17161 | 655 | lemma size_eq_Suc_imp_elem: "size M = Suc n ==> \<exists>a. a :# M" | 
| 56656 | 656 | apply (unfold size_multiset_overloaded_eq) | 
| 26178 | 657 | apply (drule setsum_SucD) | 
| 658 | apply auto | |
| 659 | done | |
| 10249 | 660 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 662 | 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 | 663 |   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 | 664 | 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 | 665 | 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 | 666 | 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 | 667 |   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 | 668 | then show ?thesis by blast | 
| 23611 | 669 | qed | 
| 15869 | 670 | |
| 26016 | 671 | |
| 672 | subsection {* Induction and case splits *}
 | |
| 10249 | 673 | |
| 18258 | 674 | theorem multiset_induct [case_names empty add, induct type: multiset]: | 
| 48009 | 675 |   assumes empty: "P {#}"
 | 
| 676 |   assumes add: "\<And>M x. P M \<Longrightarrow> P (M + {#x#})"
 | |
| 677 | shows "P M" | |
| 678 | proof (induct n \<equiv> "size M" arbitrary: M) | |
| 679 | case 0 thus "P M" by (simp add: empty) | |
| 680 | next | |
| 681 | case (Suc k) | |
| 682 |   obtain N x where "M = N + {#x#}"
 | |
| 683 | using `Suc k = size M` [symmetric] | |
| 684 | using size_eq_Suc_imp_eq_union by fast | |
| 685 | with Suc add show "P M" by simp | |
| 10249 | 686 | qed | 
| 687 | ||
| 25610 | 688 | lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = A + {#a#}"
 | 
| 26178 | 689 | by (induct M) auto | 
| 25610 | 690 | |
| 55913 | 691 | lemma multiset_cases [cases type]: | 
| 692 |   obtains (empty) "M = {#}"
 | |
| 693 |     | (add) N x where "M = N + {#x#}"
 | |
| 694 | using assms by (induct M) simp_all | |
| 25610 | 695 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 696 | 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 | 697 | 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 | 698 | |
| 26033 | 699 | 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 | 700 | apply (subst multiset_eq_iff) | 
| 26178 | 701 | apply auto | 
| 702 | done | |
| 10249 | 703 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 704 | 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 | 705 | 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 | 706 | 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 | 707 |   then have "M \<noteq> {#}" by (simp add: mset_less_empty_nonempty)
 | 
| 58425 | 708 |   then obtain M' x where "M = M' + {#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 | 709 | 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 | 710 | 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 | 711 | 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 | 712 | case (add S x T) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 713 | 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 | 714 |   have SxsubT: "S + {#x#} < T" by fact
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 715 | then have "x \<in># T" and "S < T" by (auto dest: mset_less_insertD) | 
| 58425 | 716 |   then obtain T' where T: "T = T' + {#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 | 717 | by (blast dest: multi_member_split) | 
| 58425 | 718 | 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 | 719 | 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 | 720 | 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 | 721 | 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 | 722 | 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 | 723 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 726 | |
| 58098 | 727 | text {* Well-foundedness of strict subset relation *}
 | 
| 728 | ||
| 729 | lemma wf_less_mset_rel: "wf {(M, N :: 'a multiset). M < N}"
 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 730 | 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 | 731 | 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 | 732 | 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 | 733 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 full_multiset_induct [case_names less]: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 735 | 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 | 736 | shows "P B" | 
| 58098 | 737 | apply (rule wf_less_mset_rel [THEN wf_induct]) | 
| 738 | apply (rule ih, auto) | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 739 | 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 | 740 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma multi_subset_induct [consumes 2, case_names empty add]: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 742 | 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 | 743 |   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 | 744 |   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 | 745 | 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 | 746 | proof - | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 747 | 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 | 748 | 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 | 749 | 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 | 750 |     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 | 751 | 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 | 752 | fix x F | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 753 |     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 | 754 |     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 | 755 | 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 | 756 | 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 | 757 | 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 | 758 | 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 | 759 | 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 | 760 | 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 | 761 | qed | 
| 26145 | 762 | |
| 17161 | 763 | |
| 48023 | 764 | subsection {* The fold combinator *}
 | 
| 765 | ||
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 766 | definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b"
 | 
| 48023 | 767 | where | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 768 | "fold f s M = Finite_Set.fold (\<lambda>x. f x ^^ count M x) s (set_of M)" | 
| 48023 | 769 | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 770 | lemma fold_mset_empty [simp]: | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 771 |   "fold f s {#} = s"
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 772 | by (simp add: fold_def) | 
| 48023 | 773 | |
| 774 | context comp_fun_commute | |
| 775 | begin | |
| 776 | ||
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 777 | lemma fold_mset_insert: | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 778 |   "fold f s (M + {#x#}) = f x (fold f s M)"
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 779 | proof - | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 780 | interpret mset: comp_fun_commute "\<lambda>y. f y ^^ count M y" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 781 | by (fact comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 782 |   interpret mset_union: comp_fun_commute "\<lambda>y. f y ^^ count (M + {#x#}) y"
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 783 | by (fact comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 784 | show ?thesis | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 785 | proof (cases "x \<in> set_of M") | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 786 | case False | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 787 |     then have *: "count (M + {#x#}) x = 1" by simp
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 788 |     from False have "Finite_Set.fold (\<lambda>y. f y ^^ count (M + {#x#}) y) s (set_of M) =
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 789 | Finite_Set.fold (\<lambda>y. f y ^^ count M y) s (set_of M)" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 790 | by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 791 | with False * show ?thesis | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 792 | by (simp add: fold_def del: count_union) | 
| 48023 | 793 | next | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 794 | case True | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 795 |     def N \<equiv> "set_of M - {x}"
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 796 | from N_def True have *: "set_of M = insert x N" "x \<notin> N" "finite N" by auto | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 797 |     then have "Finite_Set.fold (\<lambda>y. f y ^^ count (M + {#x#}) y) s N =
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 798 | Finite_Set.fold (\<lambda>y. f y ^^ count M y) s N" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 799 | by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 800 | with * show ?thesis by (simp add: fold_def del: count_union) simp | 
| 48023 | 801 | qed | 
| 802 | qed | |
| 803 | ||
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 804 | corollary fold_mset_single [simp]: | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 805 |   "fold f s {#x#} = f x s"
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 806 | proof - | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 807 |   have "fold f s ({#} + {#x#}) = f x s" by (simp only: fold_mset_insert) simp
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 808 | then show ?thesis by simp | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 809 | qed | 
| 48023 | 810 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 811 | lemma fold_mset_fun_left_comm: | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 812 | "f x (fold f s M) = fold f (f x s) M" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 813 | by (induct M) (simp_all add: fold_mset_insert fun_left_comm) | 
| 48023 | 814 | |
| 815 | lemma fold_mset_union [simp]: | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 816 | "fold f s (M + N) = fold f (fold f s M) N" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 817 | proof (induct M) | 
| 48023 | 818 | case empty then show ?case by simp | 
| 819 | next | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 820 | case (add M x) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 821 |   have "M + {#x#} + N = (M + N) + {#x#}"
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 822 | by (simp add: ac_simps) | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 823 | with add show ?case by (simp add: fold_mset_insert fold_mset_fun_left_comm) | 
| 48023 | 824 | qed | 
| 825 | ||
| 826 | lemma fold_mset_fusion: | |
| 827 | assumes "comp_fun_commute g" | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 828 | shows "(\<And>x y. h (g x y) = f x (h y)) \<Longrightarrow> h (fold g w A) = fold f (h w) A" (is "PROP ?P") | 
| 48023 | 829 | proof - | 
| 830 | interpret comp_fun_commute g by (fact assms) | |
| 831 | show "PROP ?P" by (induct A) auto | |
| 832 | qed | |
| 833 | ||
| 834 | end | |
| 835 | ||
| 836 | text {*
 | |
| 837 | A note on code generation: When defining some function containing a | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 838 |   subterm @{term "fold F"}, code generation is not automatic. When
 | 
| 48023 | 839 |   interpreting locale @{text left_commutative} with @{text F}, the
 | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 840 |   would be code thms for @{const fold} become thms like
 | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 841 |   @{term "fold F z {#} = z"} where @{text F} is not a pattern but
 | 
| 48023 | 842 | contains defined symbols, i.e.\ is not a code thm. Hence a separate | 
| 843 |   constant with its own code thms needs to be introduced for @{text
 | |
| 844 | F}. See the image operator below. | |
| 845 | *} | |
| 846 | ||
| 847 | ||
| 848 | subsection {* Image *}
 | |
| 849 | ||
| 850 | definition image_mset :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
 | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 851 |   "image_mset f = fold (plus o single o f) {#}"
 | 
| 48023 | 852 | |
| 49823 | 853 | lemma comp_fun_commute_mset_image: | 
| 854 | "comp_fun_commute (plus o single o f)" | |
| 855 | proof | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 856 | qed (simp add: ac_simps fun_eq_iff) | 
| 48023 | 857 | |
| 858 | lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
 | |
| 49823 | 859 | by (simp add: image_mset_def) | 
| 48023 | 860 | |
| 861 | lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}"
 | |
| 49823 | 862 | proof - | 
| 863 | interpret comp_fun_commute "plus o single o f" | |
| 864 | by (fact comp_fun_commute_mset_image) | |
| 865 | show ?thesis by (simp add: image_mset_def) | |
| 866 | qed | |
| 48023 | 867 | |
| 868 | lemma image_mset_union [simp]: | |
| 49823 | 869 | "image_mset f (M + N) = image_mset f M + image_mset f N" | 
| 870 | proof - | |
| 871 | interpret comp_fun_commute "plus o single o f" | |
| 872 | by (fact comp_fun_commute_mset_image) | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 873 | show ?thesis by (induct N) (simp_all add: image_mset_def ac_simps) | 
| 49823 | 874 | qed | 
| 875 | ||
| 876 | corollary image_mset_insert: | |
| 877 |   "image_mset f (M + {#a#}) = image_mset f M + {#f a#}"
 | |
| 878 | by simp | |
| 48023 | 879 | |
| 49823 | 880 | lemma set_of_image_mset [simp]: | 
| 881 | "set_of (image_mset f M) = image f (set_of M)" | |
| 882 | by (induct M) simp_all | |
| 48040 | 883 | |
| 49823 | 884 | lemma size_image_mset [simp]: | 
| 885 | "size (image_mset f M) = size M" | |
| 886 | by (induct M) simp_all | |
| 48023 | 887 | |
| 49823 | 888 | lemma image_mset_is_empty_iff [simp]: | 
| 889 |   "image_mset f M = {#} \<longleftrightarrow> M = {#}"
 | |
| 890 | by (cases M) auto | |
| 48023 | 891 | |
| 892 | syntax | |
| 893 | "_comprehension1_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset" | |
| 894 |       ("({#_/. _ :# _#})")
 | |
| 895 | translations | |
| 896 |   "{#e. x:#M#}" == "CONST image_mset (%x. e) M"
 | |
| 897 | ||
| 898 | syntax | |
| 899 | "_comprehension2_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset" | |
| 900 |       ("({#_/ | _ :# _./ _#})")
 | |
| 901 | translations | |
| 902 |   "{#e | x:#M. P#}" => "{#e. x :# {# x:#M. P#}#}"
 | |
| 903 | ||
| 904 | text {*
 | |
| 905 |   This allows to write not just filters like @{term "{#x:#M. x<c#}"}
 | |
| 906 |   but also images like @{term "{#x+x. x:#M #}"} and @{term [source]
 | |
| 907 |   "{#x+x|x:#M. x<c#}"}, where the latter is currently displayed as
 | |
| 908 |   @{term "{#x+x|x:#M. x<c#}"}.
 | |
| 909 | *} | |
| 910 | ||
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55417diff
changeset | 911 | functor image_mset: image_mset | 
| 48023 | 912 | proof - | 
| 913 | fix f g show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)" | |
| 914 | proof | |
| 915 | fix A | |
| 916 | show "(image_mset f \<circ> image_mset g) A = image_mset (f \<circ> g) A" | |
| 917 | by (induct A) simp_all | |
| 918 | qed | |
| 919 | show "image_mset id = id" | |
| 920 | proof | |
| 921 | fix A | |
| 922 | show "image_mset id A = id A" | |
| 923 | by (induct A) simp_all | |
| 924 | qed | |
| 925 | qed | |
| 926 | ||
| 49717 | 927 | declare image_mset.identity [simp] | 
| 928 | ||
| 48023 | 929 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 930 | subsection {* Further conversions *}
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 931 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 932 | 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 | 933 |   "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 | 934 |   "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 | 935 | |
| 37107 | 936 | lemma in_multiset_in_set: | 
| 937 | "x \<in># multiset_of xs \<longleftrightarrow> x \<in> set xs" | |
| 938 | by (induct xs) simp_all | |
| 939 | ||
| 940 | lemma count_multiset_of: | |
| 941 | "count (multiset_of xs) x = length (filter (\<lambda>y. x = y) xs)" | |
| 942 | by (induct xs) simp_all | |
| 943 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 944 | 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 | 945 | 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 | 946 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 947 | 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 | 948 | 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 | 949 | |
| 40950 | 950 | 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 | 951 | 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 | 952 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 953 | 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 | 954 | 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 | 955 | |
| 48012 | 956 | lemma size_multiset_of [simp]: "size (multiset_of xs) = length xs" | 
| 957 | by (induct xs) simp_all | |
| 958 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 959 | 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 | 960 | "multiset_of (xs @ ys) = multiset_of xs + multiset_of ys" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 961 | by (induct xs arbitrary: ys) (auto simp: ac_simps) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 962 | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 963 | lemma multiset_of_filter: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 964 |   "multiset_of (filter P xs) = {#x :# multiset_of xs. P x #}"
 | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 965 | by (induct xs) simp_all | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 966 | |
| 40950 | 967 | lemma multiset_of_rev [simp]: | 
| 968 | "multiset_of (rev xs) = multiset_of xs" | |
| 969 | by (induct xs) simp_all | |
| 970 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 971 | 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 | 972 | 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 | 973 | 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 | 974 | 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 | 975 | 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 | 976 | 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 | 977 | 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 | 978 | 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 | 979 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 980 | 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 | 981 | 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 | 982 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 983 | 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 | 984 | "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 | 985 | apply (induct x, simp, rule iffI, simp_all) | 
| 55417 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 blanchet parents: 
55129diff
changeset | 986 | apply (rename_tac 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 | 987 | 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 | 988 | 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 | 989 | 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 | 990 | 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 | 991 | 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 | 992 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 993 | 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 | 994 | "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 | 995 | 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 | 996 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 998 | "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 | 999 | (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 | 1000 | 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 | 1001 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 1003 | "(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 | 1004 | 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 | 1005 | 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 | 1006 | 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 | 1007 | [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 | 1008 | 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 | 1009 | 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 | 1010 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 1012 | "multiset_of [x\<leftarrow>xs. P x] + multiset_of [x\<leftarrow>xs. \<not>P x] = multiset_of xs" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1013 | by (induct xs) (auto simp: ac_simps) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 1015 | 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 | 1016 | "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 | 1017 | 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 | 1018 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 1020 | 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 | 1021 | 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 | 1022 | 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 | 1023 | 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 | 1024 | 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 | 1025 | |
| 36903 | 1026 | lemma multiset_of_remove1[simp]: | 
| 1027 |   "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 | 1028 | 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 | 1029 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma multiset_of_eq_length: | 
| 37107 | 1031 | assumes "multiset_of xs = multiset_of ys" | 
| 1032 | shows "length xs = length ys" | |
| 48012 | 1033 | using assms by (metis size_multiset_of) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1035 | 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 | 1036 | 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 | 1037 | shows "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) ys)" | 
| 48012 | 1038 | using assms by (metis count_multiset_of) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1039 | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1040 | lemma fold_multiset_equiv: | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1041 | assumes f: "\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x" | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1042 | and equiv: "multiset_of xs = multiset_of ys" | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1043 | shows "List.fold f xs = List.fold f ys" | 
| 46921 | 1044 | using f equiv [symmetric] | 
| 1045 | proof (induct xs arbitrary: ys) | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1046 | case Nil then show ?case by simp | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1047 | next | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1048 | case (Cons x xs) | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1049 | then have *: "set ys = set (x # xs)" by (blast dest: multiset_of_eq_setD) | 
| 58425 | 1050 | have "\<And>x y. x \<in> set ys \<Longrightarrow> y \<in> set ys \<Longrightarrow> f x \<circ> f y = f y \<circ> f x" | 
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1051 | by (rule Cons.prems(1)) (simp_all add: *) | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1052 | moreover from * have "x \<in> set ys" by simp | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1053 | ultimately have "List.fold f ys = List.fold f (remove1 x ys) \<circ> f x" by (fact fold_remove1_split) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1054 | moreover from Cons.prems have "List.fold f xs = List.fold f (remove1 x ys)" by (auto intro: Cons.hyps) | 
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1055 | ultimately show ?case by simp | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1056 | qed | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1057 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1058 | lemma multiset_of_insort [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1059 |   "multiset_of (insort x xs) = multiset_of xs + {#x#}"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1060 | by (induct xs) (simp_all add: ac_simps) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1061 | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1062 | lemma in_multiset_of: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1063 | "x \<in># multiset_of xs \<longleftrightarrow> x \<in> set xs" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1064 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1065 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1066 | lemma multiset_of_map: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1067 | "multiset_of (map f xs) = image_mset f (multiset_of xs)" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1068 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1069 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1070 | definition multiset_of_set :: "'a set \<Rightarrow> 'a multiset" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1071 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1072 |   "multiset_of_set = folding.F (\<lambda>x M. {#x#} + M) {#}"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1073 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1074 | interpretation multiset_of_set!: folding "\<lambda>x M. {#x#} + M" "{#}"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1075 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1076 |   "folding.F (\<lambda>x M. {#x#} + M) {#} = multiset_of_set"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1077 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1078 |   interpret comp_fun_commute "\<lambda>x M. {#x#} + M" by default (simp add: fun_eq_iff ac_simps)
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1079 |   show "folding (\<lambda>x M. {#x#} + M)" by default (fact comp_fun_commute)
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1080 |   from multiset_of_set_def show "folding.F (\<lambda>x M. {#x#} + M) {#} = multiset_of_set" ..
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1081 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1082 | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1083 | lemma count_multiset_of_set [simp]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1084 | "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> count (multiset_of_set A) x = 1" (is "PROP ?P") | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1085 | "\<not> finite A \<Longrightarrow> count (multiset_of_set A) x = 0" (is "PROP ?Q") | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1086 | "x \<notin> A \<Longrightarrow> count (multiset_of_set A) x = 0" (is "PROP ?R") | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1087 | proof - | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1088 |   { fix A
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1089 | assume "x \<notin> A" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1090 | have "count (multiset_of_set A) x = 0" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1091 | proof (cases "finite A") | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1092 | case False then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1093 | next | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1094 | case True from True `x \<notin> A` show ?thesis by (induct A) auto | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1095 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1096 | } note * = this | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1097 | then show "PROP ?P" "PROP ?Q" "PROP ?R" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1098 | by (auto elim!: Set.set_insert) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1099 | qed -- {* TODO: maybe define @{const multiset_of_set} also in terms of @{const Abs_multiset} *}
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1100 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1101 | context linorder | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1102 | begin | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1103 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1104 | definition sorted_list_of_multiset :: "'a multiset \<Rightarrow> 'a list" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1105 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1106 | "sorted_list_of_multiset M = fold insort [] M" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1107 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1108 | lemma sorted_list_of_multiset_empty [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1109 |   "sorted_list_of_multiset {#} = []"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1110 | by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1111 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1112 | lemma sorted_list_of_multiset_singleton [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1113 |   "sorted_list_of_multiset {#x#} = [x]"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1114 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1115 | interpret comp_fun_commute insort by (fact comp_fun_commute_insort) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1116 | show ?thesis by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1117 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1118 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1119 | lemma sorted_list_of_multiset_insert [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1120 |   "sorted_list_of_multiset (M + {#x#}) = List.insort x (sorted_list_of_multiset M)"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1121 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1122 | interpret comp_fun_commute insort by (fact comp_fun_commute_insort) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1123 | show ?thesis by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1124 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1125 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1126 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1127 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1128 | lemma multiset_of_sorted_list_of_multiset [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1129 | "multiset_of (sorted_list_of_multiset M) = M" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1130 | by (induct M) simp_all | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1131 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1132 | lemma sorted_list_of_multiset_multiset_of [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1133 | "sorted_list_of_multiset (multiset_of xs) = sort xs" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1134 | by (induct xs) simp_all | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1135 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1136 | lemma finite_set_of_multiset_of_set: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1137 | assumes "finite A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1138 | shows "set_of (multiset_of_set A) = A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1139 | using assms by (induct A) simp_all | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1140 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1141 | lemma infinite_set_of_multiset_of_set: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1142 | assumes "\<not> finite A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1143 |   shows "set_of (multiset_of_set A) = {}"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1144 | using assms by simp | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1145 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1146 | lemma set_sorted_list_of_multiset [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1147 | "set (sorted_list_of_multiset M) = set_of M" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1148 | by (induct M) (simp_all add: set_insort) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1149 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1150 | lemma sorted_list_of_multiset_of_set [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1151 | "sorted_list_of_multiset (multiset_of_set A) = sorted_list_of_set A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1152 | by (cases "finite A") (induct A rule: finite_induct, simp_all add: ac_simps) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1153 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1154 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1155 | subsection {* Big operators *}
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1156 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1157 | no_notation times (infixl "*" 70) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1158 | no_notation Groups.one ("1")
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1159 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1160 | locale comm_monoid_mset = comm_monoid | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1161 | begin | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1162 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1163 | definition F :: "'a multiset \<Rightarrow> 'a" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1164 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1165 | eq_fold: "F M = Multiset.fold f 1 M" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1166 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1167 | lemma empty [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1168 |   "F {#} = 1"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1169 | by (simp add: eq_fold) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1170 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1171 | lemma singleton [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1172 |   "F {#x#} = x"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1173 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1174 | interpret comp_fun_commute | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1175 | by default (simp add: fun_eq_iff left_commute) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1176 | show ?thesis by (simp add: eq_fold) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1177 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1178 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1179 | lemma union [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1180 | "F (M + N) = F M * F N" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1181 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1182 | interpret comp_fun_commute f | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1183 | by default (simp add: fun_eq_iff left_commute) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1184 | show ?thesis by (induct N) (simp_all add: left_commute eq_fold) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1185 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1186 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1187 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1188 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1189 | notation times (infixl "*" 70) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1190 | notation Groups.one ("1")
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1191 | |
| 54868 | 1192 | context comm_monoid_add | 
| 1193 | begin | |
| 1194 | ||
| 1195 | definition msetsum :: "'a multiset \<Rightarrow> 'a" | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1196 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1197 | "msetsum = comm_monoid_mset.F plus 0" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1198 | |
| 54868 | 1199 | sublocale msetsum!: comm_monoid_mset plus 0 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1200 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1201 | "comm_monoid_mset.F plus 0 = msetsum" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1202 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1203 | show "comm_monoid_mset plus 0" .. | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1204 | from msetsum_def show "comm_monoid_mset.F plus 0 = msetsum" .. | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1205 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1206 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1207 | lemma setsum_unfold_msetsum: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1208 | "setsum f A = msetsum (image_mset f (multiset_of_set A))" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1209 | by (cases "finite A") (induct A rule: finite_induct, simp_all) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1210 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1211 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1212 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1213 | syntax | 
| 58425 | 1214 | "_msetsum_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1215 |       ("(3SUM _:#_. _)" [0, 51, 10] 10)
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1216 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1217 | syntax (xsymbols) | 
| 58425 | 1218 | "_msetsum_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" | 
| 57518 
2f640245fc6d
refrain from auxiliary abbreviation: be more explicit to the reader in situations where syntax translation does not apply;
 haftmann parents: 
57514diff
changeset | 1219 |       ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1220 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1221 | syntax (HTML output) | 
| 58425 | 1222 | "_msetsum_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1223 |       ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1224 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1225 | translations | 
| 57518 
2f640245fc6d
refrain from auxiliary abbreviation: be more explicit to the reader in situations where syntax translation does not apply;
 haftmann parents: 
57514diff
changeset | 1226 | "SUM i :# A. b" == "CONST msetsum (CONST image_mset (\<lambda>i. b) A)" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1227 | |
| 54868 | 1228 | context comm_monoid_mult | 
| 1229 | begin | |
| 1230 | ||
| 1231 | definition msetprod :: "'a multiset \<Rightarrow> 'a" | |
| 1232 | where | |
| 1233 | "msetprod = comm_monoid_mset.F times 1" | |
| 1234 | ||
| 1235 | sublocale msetprod!: comm_monoid_mset times 1 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1236 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1237 | "comm_monoid_mset.F times 1 = msetprod" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1238 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1239 | show "comm_monoid_mset times 1" .. | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1240 | from msetprod_def show "comm_monoid_mset.F times 1 = msetprod" .. | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1241 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1242 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1243 | lemma msetprod_empty: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1244 |   "msetprod {#} = 1"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1245 | by (fact msetprod.empty) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1246 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1247 | lemma msetprod_singleton: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1248 |   "msetprod {#x#} = x"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1249 | by (fact msetprod.singleton) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1250 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1251 | lemma msetprod_Un: | 
| 58425 | 1252 | "msetprod (A + B) = msetprod A * msetprod B" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1253 | by (fact msetprod.union) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1254 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1255 | lemma setprod_unfold_msetprod: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1256 | "setprod f A = msetprod (image_mset f (multiset_of_set A))" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1257 | by (cases "finite A") (induct A rule: finite_induct, simp_all) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1258 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1259 | lemma msetprod_multiplicity: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1260 | "msetprod M = setprod (\<lambda>x. x ^ count M x) (set_of M)" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1261 | by (simp add: Multiset.fold_def setprod.eq_fold msetprod.eq_fold funpow_times_power comp_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1262 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1263 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1264 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1265 | syntax | 
| 58425 | 1266 | "_msetprod_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1267 |       ("(3PROD _:#_. _)" [0, 51, 10] 10)
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1268 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1269 | syntax (xsymbols) | 
| 58425 | 1270 | "_msetprod_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1271 |       ("(3\<Prod>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1272 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1273 | syntax (HTML output) | 
| 58425 | 1274 | "_msetprod_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1275 |       ("(3\<Prod>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1276 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1277 | translations | 
| 57518 
2f640245fc6d
refrain from auxiliary abbreviation: be more explicit to the reader in situations where syntax translation does not apply;
 haftmann parents: 
57514diff
changeset | 1278 | "PROD i :# A. b" == "CONST msetprod (CONST image_mset (\<lambda>i. b) A)" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1279 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1280 | lemma (in comm_semiring_1) dvd_msetprod: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1281 | assumes "x \<in># A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1282 | shows "x dvd msetprod A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1283 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1284 |   from assms have "A = (A - {#x#}) + {#x#}" by simp
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1285 |   then obtain B where "A = B + {#x#}" ..
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1286 | then show ?thesis by simp | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1287 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1288 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1289 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1290 | subsection {* Cardinality *}
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1291 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1292 | definition mcard :: "'a multiset \<Rightarrow> nat" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1293 | where | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1294 | "mcard = msetsum \<circ> image_mset (\<lambda>_. 1)" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1295 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1296 | lemma mcard_empty [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1297 |   "mcard {#} = 0"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1298 | by (simp add: mcard_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1299 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1300 | lemma mcard_singleton [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1301 |   "mcard {#a#} = Suc 0"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1302 | by (simp add: mcard_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1303 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1304 | lemma mcard_plus [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1305 | "mcard (M + N) = mcard M + mcard N" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1306 | by (simp add: mcard_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1307 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1308 | lemma mcard_empty_iff [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1309 |   "mcard M = 0 \<longleftrightarrow> M = {#}"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1310 | by (induct M) simp_all | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1311 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1312 | lemma mcard_unfold_setsum: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1313 | "mcard M = setsum (count M) (set_of M)" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1314 | proof (induct M) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1315 | case empty then show ?case by simp | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1316 | next | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1317 | case (add M x) then show ?case | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1318 | by (cases "x \<in> set_of M") | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1319 | (simp_all del: mem_set_of_iff add: setsum.distrib setsum.delta' insert_absorb, simp) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1320 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1321 | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1322 | lemma size_eq_mcard: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1323 | "size = mcard" | 
| 56656 | 1324 | by (simp add: fun_eq_iff size_multiset_overloaded_eq mcard_unfold_setsum) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1325 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1326 | lemma mcard_multiset_of: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1327 | "mcard (multiset_of xs) = length xs" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1328 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1329 | |
| 58035 | 1330 | lemma mcard_mono: assumes "A \<le> B" | 
| 1331 | shows "mcard A \<le> mcard B" | |
| 1332 | proof - | |
| 1333 | from assms[unfolded mset_le_exists_conv] | |
| 1334 | obtain C where B: "B = A + C" by auto | |
| 1335 | show ?thesis unfolding B by (induct C, auto) | |
| 1336 | qed | |
| 1337 | ||
| 1338 | lemma mcard_filter_lesseq[simp]: "mcard (Multiset.filter f M) \<le> mcard M" | |
| 1339 | by (rule mcard_mono[OF multiset_filter_subset]) | |
| 1340 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1341 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1342 | subsection {* Alternative representations *}
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1343 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1344 | subsubsection {* Lists *}
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1345 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1346 | context linorder | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1347 | begin | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1348 | |
| 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 | 1349 | 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 | 1350 |   "multiset_of (insort_key k x xs) = {#x#} + multiset_of xs"
 | 
| 37107 | 1351 | by (induct xs) (simp_all add: ac_simps) | 
| 58425 | 1352 | |
| 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 | 1353 | 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 | 1354 | "multiset_of (sort_key k xs) = multiset_of xs" | 
| 37107 | 1355 | by (induct xs) (simp_all add: ac_simps) | 
| 1356 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1357 | 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 | 1358 | 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 | 1359 |   @{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 | 1360 | *} | 
| 37074 | 1361 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1362 | lemma properties_for_sort_key: | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1363 | assumes "multiset_of ys = multiset_of xs" | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1364 | 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 | 1365 | and "sorted (map f ys)" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1366 | shows "sort_key f xs = ys" | 
| 46921 | 1367 | using assms | 
| 1368 | 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 | 1369 | 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 | 1370 | 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 | 1371 | case (Cons x xs) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1372 | from Cons.prems(2) have | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1373 | "\<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 | 1374 | by (simp add: filter_remove1) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1375 | 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 | 1376 | 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 | 1377 | 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 | 1378 | 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 | 1379 | 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 | 1380 | 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 | 1381 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1382 | lemma properties_for_sort: | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1383 | 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 | 1384 | and "sorted ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1385 | shows "sort xs = ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1386 | 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 | 1387 | 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 | 1388 | 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 | 1389 | 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 | 1390 | 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 | 1391 | 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 | 1392 | by simp | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1393 | 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 | 1394 | 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 | 1395 | qed | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1396 | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1397 | lemma sort_key_by_quicksort: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1398 | "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 | 1399 | @ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1400 | @ 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 | 1401 | proof (rule properties_for_sort_key) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1402 | show "multiset_of ?rhs = multiset_of ?lhs" | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1403 | by (rule multiset_eqI) (auto simp add: multiset_of_filter) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1404 | next | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1405 | show "sorted (map f ?rhs)" | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1406 | by (auto simp add: sorted_append intro: sorted_map_same) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1407 | next | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1408 | fix l | 
| 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1409 | assume "l \<in> set ?rhs" | 
| 40346 | 1410 | let ?pivot = "f (xs ! (length xs div 2))" | 
| 1411 | have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto | |
| 40306 | 1412 | 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 | 1413 | unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) | 
| 40346 | 1414 | with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp | 
| 1415 | 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 | |
| 1416 | then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] = | |
| 1417 | [x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp | |
| 1418 | note *** = this [of "op <"] this [of "op >"] this [of "op ="] | |
| 40306 | 1419 | 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 | 1420 | proof (cases "f l" ?pivot rule: linorder_cases) | 
| 46730 | 1421 | case less | 
| 1422 | then have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto | |
| 1423 | with less show ?thesis | |
| 40346 | 1424 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1425 | next | 
| 40306 | 1426 | case equal then show ?thesis | 
| 40346 | 1427 | by (simp add: * less_le) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1428 | next | 
| 46730 | 1429 | case greater | 
| 1430 | then have "f l \<noteq> ?pivot" and "\<not> f l < ?pivot" by auto | |
| 1431 | with greater show ?thesis | |
| 40346 | 1432 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40306 | 1433 | qed | 
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1434 | qed | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1435 | |
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1436 | lemma sort_by_quicksort: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1437 | "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 | 1438 | @ [x\<leftarrow>xs. x = xs ! (length xs div 2)] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1439 | @ 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 | 1440 | 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 | 1441 | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1442 | text {* A stable parametrized quicksort *}
 | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1443 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1444 | 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 | 1445 | "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 | 1446 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1447 | lemma part_code [code]: | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1448 | "part f pivot [] = ([], [], [])" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1449 | "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 | 1450 | if x' < pivot then (x # lts, eqs, gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1451 | else if x' > pivot then (lts, eqs, x # gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1452 | else (lts, x # eqs, gts))" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1453 | by (auto simp add: part_def Let_def split_def) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1454 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1455 | lemma sort_key_by_quicksort_code [code]: | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1456 | "sort_key f xs = (case xs of [] \<Rightarrow> [] | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1457 | | [x] \<Rightarrow> xs | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1458 | | [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 | 1459 | | _ \<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 | 1460 | in sort_key f lts @ eqs @ sort_key f gts))" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1461 | proof (cases xs) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1462 | case Nil then show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1463 | next | 
| 46921 | 1464 | case (Cons _ ys) note hyps = Cons show ?thesis | 
| 1465 | proof (cases ys) | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1466 | case Nil with hyps show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1467 | next | 
| 46921 | 1468 | case (Cons _ zs) note hyps = hyps Cons show ?thesis | 
| 1469 | proof (cases zs) | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1470 | case Nil with hyps show ?thesis by auto | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1471 | next | 
| 58425 | 1472 | case Cons | 
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1473 | from sort_key_by_quicksort [of f xs] | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1474 | 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 | 1475 | in sort_key f lts @ eqs @ sort_key f gts)" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1476 | 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 | 1477 | with hyps Cons show ?thesis by (simp only: list.cases) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1478 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1479 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1480 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1481 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1482 | end | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1483 | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1484 | hide_const (open) part | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1485 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1486 | 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 | 1487 | 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 | 1488 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1489 | 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 | 1490 |   "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 | 1491 | 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 | 1492 | 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 | 1493 | 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 | 1494 | 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 | 1495 | 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 | 1496 | 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 | 1497 | 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 | 1498 | 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 | 1499 | 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 | 1500 | 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 | 1501 | apply simp | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1502 | apply (subst add.assoc) | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1503 |       apply (subst add.commute [of "{#v#}" "{#x#}"])
 | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1504 | apply (subst add.assoc [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 | 1505 | 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 | 1506 | 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 | 1507 | 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 | 1508 | 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 | 1509 | 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 | 1510 | 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 | 1511 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1512 | 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 | 1513 | "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 | 1514 | 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 | 1515 | 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 | 1516 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1517 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | subsection {* The multiset order *}
 | 
| 10249 | 1519 | |
| 1520 | subsubsection {* Well-foundedness *}
 | |
| 1521 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1522 | definition mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
 | 
| 37765 | 1523 |   "mult1 r = {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
 | 
| 23751 | 1524 | (\<forall>b. b :# K --> (b, a) \<in> r)}" | 
| 10249 | 1525 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1526 | definition mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
 | 
| 37765 | 1527 | "mult r = (mult1 r)\<^sup>+" | 
| 10249 | 1528 | |
| 23751 | 1529 | lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
 | 
| 26178 | 1530 | by (simp add: mult1_def) | 
| 10249 | 1531 | |
| 23751 | 1532 | lemma less_add: "(N, M0 + {#a#}) \<in> mult1 r ==>
 | 
| 1533 |     (\<exists>M. (M, M0) \<in> mult1 r \<and> N = M + {#a#}) \<or>
 | |
| 1534 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K)" | |
| 19582 | 1535 | (is "_ \<Longrightarrow> ?case1 (mult1 r) \<or> ?case2") | 
| 10249 | 1536 | proof (unfold mult1_def) | 
| 23751 | 1537 | let ?r = "\<lambda>K a. \<forall>b. b :# K --> (b, a) \<in> r" | 
| 11464 | 1538 |   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
 | 
| 23751 | 1539 |   let ?case1 = "?case1 {(N, M). ?R N M}"
 | 
| 10249 | 1540 | |
| 23751 | 1541 |   assume "(N, M0 + {#a#}) \<in> {(N, M). ?R N M}"
 | 
| 18258 | 1542 | then have "\<exists>a' M0' K. | 
| 11464 | 1543 |       M0 + {#a#} = M0' + {#a'#} \<and> N = M0' + K \<and> ?r K a'" by simp
 | 
| 18258 | 1544 | then show "?case1 \<or> ?case2" | 
| 10249 | 1545 | proof (elim exE conjE) | 
| 1546 | fix a' M0' K | |
| 1547 | assume N: "N = M0' + K" and r: "?r K a'" | |
| 1548 |     assume "M0 + {#a#} = M0' + {#a'#}"
 | |
| 18258 | 1549 | then have "M0 = M0' \<and> a = a' \<or> | 
| 11464 | 1550 |         (\<exists>K'. M0 = K' + {#a'#} \<and> M0' = K' + {#a#})"
 | 
| 10249 | 1551 | by (simp only: add_eq_conv_ex) | 
| 18258 | 1552 | then show ?thesis | 
| 10249 | 1553 | proof (elim disjE conjE exE) | 
| 1554 | assume "M0 = M0'" "a = a'" | |
| 11464 | 1555 | with N r have "?r K a \<and> N = M0 + K" by simp | 
| 18258 | 1556 | then have ?case2 .. then show ?thesis .. | 
| 10249 | 1557 | next | 
| 1558 | fix K' | |
| 1559 |       assume "M0' = K' + {#a#}"
 | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1560 |       with N have n: "N = K' + K + {#a#}" by (simp add: ac_simps)
 | 
| 10249 | 1561 | |
| 1562 |       assume "M0 = K' + {#a'#}"
 | |
| 1563 | with r have "?R (K' + K) M0" by blast | |
| 18258 | 1564 | with n have ?case1 by simp then show ?thesis .. | 
| 10249 | 1565 | qed | 
| 1566 | qed | |
| 1567 | qed | |
| 1568 | ||
| 54295 | 1569 | lemma all_accessible: "wf r ==> \<forall>M. M \<in> Wellfounded.acc (mult1 r)" | 
| 10249 | 1570 | proof | 
| 1571 | let ?R = "mult1 r" | |
| 54295 | 1572 | let ?W = "Wellfounded.acc ?R" | 
| 10249 | 1573 |   {
 | 
| 1574 | fix M M0 a | |
| 23751 | 1575 | assume M0: "M0 \<in> ?W" | 
| 1576 |       and wf_hyp: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | |
| 1577 |       and acc_hyp: "\<forall>M. (M, M0) \<in> ?R --> M + {#a#} \<in> ?W"
 | |
| 1578 |     have "M0 + {#a#} \<in> ?W"
 | |
| 1579 |     proof (rule accI [of "M0 + {#a#}"])
 | |
| 10249 | 1580 | fix N | 
| 23751 | 1581 |       assume "(N, M0 + {#a#}) \<in> ?R"
 | 
| 1582 |       then have "((\<exists>M. (M, M0) \<in> ?R \<and> N = M + {#a#}) \<or>
 | |
| 1583 | (\<exists>K. (\<forall>b. b :# K --> (b, a) \<in> r) \<and> N = M0 + K))" | |
| 10249 | 1584 | by (rule less_add) | 
| 23751 | 1585 | then show "N \<in> ?W" | 
| 10249 | 1586 | proof (elim exE disjE conjE) | 
| 23751 | 1587 |         fix M assume "(M, M0) \<in> ?R" and N: "N = M + {#a#}"
 | 
| 1588 |         from acc_hyp have "(M, M0) \<in> ?R --> M + {#a#} \<in> ?W" ..
 | |
| 1589 |         from this and `(M, M0) \<in> ?R` have "M + {#a#} \<in> ?W" ..
 | |
| 1590 | then show "N \<in> ?W" by (simp only: N) | |
| 10249 | 1591 | next | 
| 1592 | fix K | |
| 1593 | assume N: "N = M0 + K" | |
| 23751 | 1594 | assume "\<forall>b. b :# K --> (b, a) \<in> r" | 
| 1595 | then have "M0 + K \<in> ?W" | |
| 10249 | 1596 | proof (induct K) | 
| 18730 | 1597 | case empty | 
| 23751 | 1598 |           from M0 show "M0 + {#} \<in> ?W" by simp
 | 
| 18730 | 1599 | next | 
| 1600 | case (add K x) | |
| 23751 | 1601 | from add.prems have "(x, a) \<in> r" by simp | 
| 1602 |           with wf_hyp have "\<forall>M \<in> ?W. M + {#x#} \<in> ?W" by blast
 | |
| 1603 | moreover from add have "M0 + K \<in> ?W" by simp | |
| 1604 |           ultimately have "(M0 + K) + {#x#} \<in> ?W" ..
 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1605 |           then show "M0 + (K + {#x#}) \<in> ?W" by (simp only: add.assoc)
 | 
| 10249 | 1606 | qed | 
| 23751 | 1607 | then show "N \<in> ?W" by (simp only: N) | 
| 10249 | 1608 | qed | 
| 1609 | qed | |
| 1610 | } note tedious_reasoning = this | |
| 1611 | ||
| 23751 | 1612 | assume wf: "wf r" | 
| 10249 | 1613 | fix M | 
| 23751 | 1614 | show "M \<in> ?W" | 
| 10249 | 1615 | proof (induct M) | 
| 23751 | 1616 |     show "{#} \<in> ?W"
 | 
| 10249 | 1617 | proof (rule accI) | 
| 23751 | 1618 |       fix b assume "(b, {#}) \<in> ?R"
 | 
| 1619 | with not_less_empty show "b \<in> ?W" by contradiction | |
| 10249 | 1620 | qed | 
| 1621 | ||
| 23751 | 1622 | fix M a assume "M \<in> ?W" | 
| 1623 |     from wf have "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | |
| 10249 | 1624 | proof induct | 
| 1625 | fix a | |
| 23751 | 1626 |       assume r: "!!b. (b, a) \<in> r ==> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | 
| 1627 |       show "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | |
| 10249 | 1628 | proof | 
| 23751 | 1629 | fix M assume "M \<in> ?W" | 
| 1630 |         then show "M + {#a#} \<in> ?W"
 | |
| 23373 | 1631 | by (rule acc_induct) (rule tedious_reasoning [OF _ r]) | 
| 10249 | 1632 | qed | 
| 1633 | qed | |
| 23751 | 1634 |     from this and `M \<in> ?W` show "M + {#a#} \<in> ?W" ..
 | 
| 10249 | 1635 | qed | 
| 1636 | qed | |
| 1637 | ||
| 23751 | 1638 | theorem wf_mult1: "wf r ==> wf (mult1 r)" | 
| 26178 | 1639 | by (rule acc_wfI) (rule all_accessible) | 
| 10249 | 1640 | |
| 23751 | 1641 | theorem wf_mult: "wf r ==> wf (mult r)" | 
| 26178 | 1642 | unfolding mult_def by (rule wf_trancl) (rule wf_mult1) | 
| 10249 | 1643 | |
| 1644 | ||
| 1645 | subsubsection {* Closure-free presentation *}
 | |
| 1646 | ||
| 1647 | text {* One direction. *}
 | |
| 1648 | ||
| 1649 | lemma mult_implies_one_step: | |
| 23751 | 1650 | "trans r ==> (M, N) \<in> mult r ==> | 
| 11464 | 1651 |     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
 | 
| 23751 | 1652 | (\<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r)" | 
| 26178 | 1653 | apply (unfold mult_def mult1_def set_of_def) | 
| 1654 | apply (erule converse_trancl_induct, clarify) | |
| 1655 | apply (rule_tac x = M0 in exI, simp, clarify) | |
| 1656 | apply (case_tac "a :# K") | |
| 1657 | apply (rule_tac x = I in exI) | |
| 1658 | apply (simp (no_asm)) | |
| 1659 |  apply (rule_tac x = "(K - {#a#}) + Ka" in exI)
 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1660 | apply (simp (no_asm_simp) add: add.assoc [symmetric]) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
57418diff
changeset | 1661 |  apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="?S + ?T" in arg_cong)
 | 
| 26178 | 1662 | apply (simp add: diff_union_single_conv) | 
| 1663 | apply (simp (no_asm_use) add: trans_def) | |
| 1664 | apply blast | |
| 1665 | apply (subgoal_tac "a :# I") | |
| 1666 |  apply (rule_tac x = "I - {#a#}" in exI)
 | |
| 1667 |  apply (rule_tac x = "J + {#a#}" in exI)
 | |
| 1668 | apply (rule_tac x = "K + Ka" in exI) | |
| 1669 | 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 | 1670 | apply (simp add: multiset_eq_iff split: nat_diff_split) | 
| 26178 | 1671 | apply (rule conjI) | 
| 57492 
74bf65a1910a
Hypsubst preserves equality hypotheses
 Thomas Sewell <thomas.sewell@nicta.com.au> parents: 
57418diff
changeset | 1672 |   apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="?S + ?T" 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 | 1673 | apply (simp add: multiset_eq_iff split: nat_diff_split) | 
| 26178 | 1674 | apply (simp (no_asm_use) add: trans_def) | 
| 1675 | apply blast | |
| 1676 | apply (subgoal_tac "a :# (M0 + {#a#})")
 | |
| 1677 | apply simp | |
| 1678 | apply (simp (no_asm)) | |
| 1679 | done | |
| 10249 | 1680 | |
| 1681 | lemma one_step_implies_mult_aux: | |
| 23751 | 1682 | "trans r ==> | 
| 1683 |     \<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))
 | |
| 1684 | --> (I + K, I + J) \<in> mult r" | |
| 26178 | 1685 | apply (induct_tac n, auto) | 
| 1686 | apply (frule size_eq_Suc_imp_eq_union, clarify) | |
| 1687 | apply (rename_tac "J'", simp) | |
| 1688 | apply (erule notE, auto) | |
| 1689 | apply (case_tac "J' = {#}")
 | |
| 1690 | apply (simp add: mult_def) | |
| 1691 | apply (rule r_into_trancl) | |
| 1692 | apply (simp add: mult1_def set_of_def, blast) | |
| 1693 | txt {* Now we know @{term "J' \<noteq> {#}"}. *}
 | |
| 1694 | apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition) | |
| 1695 | apply (erule_tac P = "\<forall>k \<in> set_of K. ?P k" in rev_mp) | |
| 1696 | apply (erule ssubst) | |
| 1697 | apply (simp add: Ball_def, auto) | |
| 1698 | apply (subgoal_tac | |
| 1699 |   "((I + {# x :# K. (x, a) \<in> r #}) + {# x :# K. (x, a) \<notin> r #},
 | |
| 1700 |     (I + {# x :# K. (x, a) \<in> r #}) + J') \<in> mult r")
 | |
| 1701 | prefer 2 | |
| 1702 | apply force | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1703 | apply (simp (no_asm_use) add: add.assoc [symmetric] mult_def) | 
| 26178 | 1704 | apply (erule trancl_trans) | 
| 1705 | apply (rule r_into_trancl) | |
| 1706 | apply (simp add: mult1_def set_of_def) | |
| 1707 | apply (rule_tac x = a in exI) | |
| 1708 | apply (rule_tac x = "I + J'" in exI) | |
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1709 | apply (simp add: ac_simps) | 
| 26178 | 1710 | done | 
| 10249 | 1711 | |
| 17161 | 1712 | lemma one_step_implies_mult: | 
| 23751 | 1713 |   "trans r ==> J \<noteq> {#} ==> \<forall>k \<in> set_of K. \<exists>j \<in> set_of J. (k, j) \<in> r
 | 
| 1714 | ==> (I + K, I + J) \<in> mult r" | |
| 26178 | 1715 | using one_step_implies_mult_aux by blast | 
| 10249 | 1716 | |
| 1717 | ||
| 1718 | subsubsection {* Partial-order properties *}
 | |
| 1719 | ||
| 35273 | 1720 | definition less_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<#" 50) where | 
| 1721 |   "M' <# M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
 | |
| 10249 | 1722 | |
| 35273 | 1723 | definition le_multiset :: "'a\<Colon>order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "<=#" 50) where | 
| 1724 | "M' <=# M \<longleftrightarrow> M' <# M \<or> M' = M" | |
| 1725 | ||
| 35308 | 1726 | notation (xsymbols) less_multiset (infix "\<subset>#" 50) | 
| 1727 | notation (xsymbols) le_multiset (infix "\<subseteq>#" 50) | |
| 10249 | 1728 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1729 | interpretation multiset_order: order le_multiset less_multiset | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1730 | proof - | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1731 | have irrefl: "\<And>M :: 'a multiset. \<not> M \<subset># M" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1732 | proof | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1733 | fix M :: "'a multiset" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1734 | assume "M \<subset># M" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1735 |     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 | 1736 |     have "trans {(x'::'a, x). x' < x}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1737 | by (rule transI) simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1738 | moreover note MM | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1739 | 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 | 1740 |       \<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 | 1741 | by (rule mult_implies_one_step) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1742 | 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 | 1743 |       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 | 1744 |     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 | 1745 | have "finite (set_of K)" by simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1746 | moreover note aux2 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1747 |     ultimately have "set_of K = {}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1748 | by (induct rule: finite_induct) (auto intro: order_less_trans) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1749 | with aux1 show False by simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1750 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1751 | 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 | 1752 | unfolding less_multiset_def mult_def by (blast intro: trancl_trans) | 
| 46921 | 1753 | show "class.order (le_multiset :: 'a multiset \<Rightarrow> _) less_multiset" | 
| 1754 | by default (auto simp add: le_multiset_def irrefl dest: trans) | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1755 | qed | 
| 10249 | 1756 | |
| 46730 | 1757 | lemma mult_less_irrefl [elim!]: "M \<subset># (M::'a::order multiset) ==> R" | 
| 1758 | by simp | |
| 26567 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 1759 | |
| 10249 | 1760 | |
| 1761 | subsubsection {* Monotonicity of multiset union *}
 | |
| 1762 | ||
| 46730 | 1763 | lemma mult1_union: "(B, D) \<in> mult1 r ==> (C + B, C + D) \<in> mult1 r" | 
| 26178 | 1764 | apply (unfold mult1_def) | 
| 1765 | apply auto | |
| 1766 | apply (rule_tac x = a in exI) | |
| 1767 | apply (rule_tac x = "C + M0" in exI) | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1768 | apply (simp add: add.assoc) | 
| 26178 | 1769 | done | 
| 10249 | 1770 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1771 | lemma union_less_mono2: "B \<subset># D ==> C + B \<subset># C + (D::'a::order multiset)" | 
| 26178 | 1772 | apply (unfold less_multiset_def mult_def) | 
| 1773 | apply (erule trancl_induct) | |
| 40249 
cd404ecb9107
Remove unnecessary premise of mult1_union
 Lars Noschinski <noschinl@in.tum.de> parents: 
39533diff
changeset | 1774 | apply (blast intro: mult1_union) | 
| 
cd404ecb9107
Remove unnecessary premise of mult1_union
 Lars Noschinski <noschinl@in.tum.de> parents: 
39533diff
changeset | 1775 | apply (blast intro: mult1_union trancl_trans) | 
| 26178 | 1776 | done | 
| 10249 | 1777 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1778 | lemma union_less_mono1: "B \<subset># D ==> B + C \<subset># D + (C::'a::order multiset)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1779 | apply (subst add.commute [of B C]) | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1780 | apply (subst add.commute [of D C]) | 
| 26178 | 1781 | apply (erule union_less_mono2) | 
| 1782 | done | |
| 10249 | 1783 | |
| 17161 | 1784 | lemma union_less_mono: | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1785 | "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 | 1786 | by (blast intro!: union_less_mono1 union_less_mono2 multiset_order.less_trans) | 
| 10249 | 1787 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1788 | interpretation multiset_order: ordered_ab_semigroup_add plus le_multiset less_multiset | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1789 | proof | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1790 | qed (auto simp add: le_multiset_def intro: union_less_mono2) | 
| 26145 | 1791 | |
| 15072 | 1792 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1793 | subsection {* Termination proofs with multiset orders *}
 | 
| 
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 | 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 | 1796 |   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 | 1797 |   and multi_member_last: "x \<in># {# x #}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1798 | by auto | 
| 
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 | definition "ms_strict = mult pair_less" | 
| 37765 | 1801 | 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 | 1802 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1803 | 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 | 1804 | 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 | 1805 | 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 | 1806 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1807 | lemma smsI: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1808 | "(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 | 1809 | unfolding ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1810 | 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 | 1811 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1812 | lemma wmsI: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1813 |   "(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 | 1814 | \<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 | 1815 | unfolding ms_weak_def ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1816 | 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 | 1817 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1818 | inductive pw_leq | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1819 | where | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1820 |   pw_leq_empty: "pw_leq {#} {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1821 | | 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 | 1822 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1823 | lemma pw_leq_lstep: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1824 |   "(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 | 1825 | 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 | 1826 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1827 | lemma pw_leq_split: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1828 | assumes "pw_leq X Y" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1829 |   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 | 1830 | using assms | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1831 | proof (induct) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1832 | 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 | 1833 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1834 | 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 | 1835 | then obtain A B Z where | 
| 58425 | 1836 | [simp]: "X = A + Z" "Y = B + Z" | 
| 1837 |       and 1[simp]: "(set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#})"
 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1838 | by auto | 
| 58425 | 1839 | from pw_leq_step have "x = y \<or> (x, y) \<in> pair_less" | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1840 | unfolding pair_leq_def by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1841 | thus ?case | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1842 | proof | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1843 | assume [simp]: "x = y" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1844 | have | 
| 58425 | 1845 |       "{#x#} + X = A + ({#y#}+Z)
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1846 |       \<and> {#y#} + Y = B + ({#y#}+Z)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1847 |       \<and> ((set_of A, set_of B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1848 | by (auto simp: ac_simps) | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1849 | thus ?case by (intro exI) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1850 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1851 | assume A: "(x, y) \<in> pair_less" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1852 |     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 | 1853 |     have "{#x#} + X = ?A' + Z"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1854 |       "{#y#} + Y = ?B' + Z"
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1855 | by (auto simp add: ac_simps) | 
| 58425 | 1856 | moreover have | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1857 | "(set_of ?A', set_of ?B') \<in> max_strict" | 
| 58425 | 1858 | using 1 A unfolding max_strict_def | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1859 | by (auto elim!: max_ext.cases) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1860 | ultimately show ?thesis by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1861 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1862 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1863 | |
| 58425 | 1864 | lemma | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1865 | assumes pwleq: "pw_leq Z Z'" | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1866 | 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 | 1867 | 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 | 1868 |   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 | 1869 | proof - | 
| 58425 | 1870 | from pw_leq_split[OF pwleq] | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1871 | obtain A' B' Z'' | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1872 | 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 | 1873 |     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 | 1874 | by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1875 |   {
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1876 | 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 | 1877 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1878 | 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 | 1879 | proof | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1880 | 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 | 1881 | 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 | 1882 | by (auto simp: max_strict_def intro: max_ext_additive) | 
| 58425 | 1883 | thus ?thesis by (rule smsI) | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1884 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1885 |       assume [simp]: "A' = {#} \<and> B' = {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1886 | 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 | 1887 | qed | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1888 | thus "(Z + A, Z' + B) \<in> ms_strict" by (simp add:ac_simps) | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1889 | 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 | 1890 | } | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1891 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1892 | have "(Z'' + A', Z'' + B') \<in> ms_weak" by (rule wmsI) | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1893 |   thus "(Z + {#}, Z' + {#}) \<in> ms_weak" by (simp add:ac_simps)
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1894 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1895 | |
| 39301 | 1896 | 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 | 1897 | and nonempty_plus: "{# x #} + rs \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1898 | and nonempty_single: "{# x #} \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1899 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1900 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1901 | setup {*
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1902 | let | 
| 35402 | 1903 |   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 | 1904 | |
| 35402 | 1905 |   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 | 1906 |     | 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 | 1907 | | mk_mset T (x :: xs) = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1908 |           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 | 1909 | 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 | 1910 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1911 | fun mset_member_tac m i = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1912 | (if m <= 0 then | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1913 |            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 | 1914 | else | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1915 |            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 | 1916 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1917 | val mset_nonempty_tac = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1918 |       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 | 1919 | |
| 59625 | 1920 | fun regroup_munion_conv ctxt = | 
| 1921 |     Function_Lib.regroup_conv ctxt @{const_abbrev Mempty} @{const_name plus}
 | |
| 1922 |       (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral}))
 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1923 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1924 | fun unfold_pwleq_tac i = | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1925 |     (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 | 1926 |       ORELSE (rtac @{thm pw_leq_lstep} i)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1927 |       ORELSE (rtac @{thm pw_leq_empty} i)
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1928 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1929 |   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 | 1930 |                       @{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 | 1931 | in | 
| 58425 | 1932 | ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1933 |   {
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1934 | 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 | 1935 | 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 | 1936 | 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 | 1937 |     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 | 1938 |     reduction_pair= @{thm ms_reduction_pair}
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1939 | }) | 
| 10249 | 1940 | end | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1941 | *} | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 1942 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1943 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1944 | 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 | 1945 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1946 | 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 | 1947 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1948 | lemma union_commute: "M + N = N + (M::'a multiset)" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1949 | by (fact add.commute) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1950 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1951 | lemma union_assoc: "(M + N) + K = M + (N + (K::'a multiset))" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1952 | by (fact add.assoc) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1953 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1954 | lemma union_lcomm: "M + (N + K) = N + (M + (K::'a multiset))" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1955 | by (fact add.left_commute) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1956 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1957 | 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 | 1958 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1959 | 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 | 1960 | 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 | 1961 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1962 | 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 | 1963 | 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 | 1964 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1965 | lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y" | 
| 59557 | 1966 | by (fact add_left_imp_eq) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1967 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1968 | 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 | 1969 | by (fact order_less_trans) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1970 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1971 | lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1972 | by (fact inf.commute) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1973 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1974 | 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 | 1975 | by (fact inf.assoc [symmetric]) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1976 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1977 | 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 | 1978 | by (fact inf.left_commute) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1979 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1980 | lemmas multiset_inter_ac = | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1981 | multiset_inter_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1982 | multiset_inter_assoc | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1983 | multiset_inter_left_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1984 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1985 | lemma mult_less_not_refl: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1986 | "\<not> M \<subset># (M::'a::order multiset)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1987 | by (fact multiset_order.less_irrefl) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1988 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1989 | lemma mult_less_trans: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1990 | "K \<subset># M ==> M \<subset># N ==> K \<subset># (N::'a::order multiset)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1991 | by (fact multiset_order.less_trans) | 
| 58425 | 1992 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1993 | lemma mult_less_not_sym: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1994 | "M \<subset># N ==> \<not> N \<subset># (M::'a::order multiset)" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1995 | by (fact multiset_order.less_not_sym) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1996 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1997 | lemma mult_less_asym: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1998 | "M \<subset># N ==> (\<not> P ==> N \<subset># (M::'a::order multiset)) ==> P" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1999 | 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 | 2000 | |
| 35712 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2001 | ML {*
 | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2002 | 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 | 2003 | (Const _ $ t') = | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2004 | let | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2005 | 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 | 2006 | 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 | 2007 | ||> 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 | 2008 | 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 | 2009 | 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 | 2010 | 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 | 2011 | | 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 | 2012 | in | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2013 | case maps elems_for (all_values elem_T) @ | 
| 37261 | 2014 | (if maybe_opt then [Const (Nitpick_Model.unrep (), elem_T)] | 
| 2015 | else []) of | |
| 35712 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2016 |         [] => 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 | 2017 | | 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 | 2018 |                          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 | 2019 | $ t1 $ t2) | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2020 |                      (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 | 2021 | 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 | 2022 | end | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2023 | | 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 | 2024 | *} | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2025 | |
| 38287 | 2026 | declaration {*
 | 
| 2027 | Nitpick_Model.register_term_postprocessor @{typ "'a multiset"}
 | |
| 38242 | 2028 | multiset_postproc | 
| 35712 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2029 | *} | 
| 
77aa29bf14ee
added a mechanism to Nitpick to support custom rendering of terms, and used it for multisets
 blanchet parents: 
35402diff
changeset | 2030 | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 2031 | hide_const (open) fold | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 2032 | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2033 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2034 | subsection {* Naive implementation using lists *}
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2035 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2036 | code_datatype multiset_of | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2037 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2038 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2039 |   "{#} = multiset_of []"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2040 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2041 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2042 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2043 |   "{#x#} = multiset_of [x]"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2044 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2045 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2046 | lemma union_code [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2047 | "multiset_of xs + multiset_of ys = multiset_of (xs @ ys)" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2048 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2049 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2050 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2051 | "image_mset f (multiset_of xs) = multiset_of (map f xs)" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2052 | by (simp add: multiset_of_map) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2053 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2054 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2055 | "Multiset.filter f (multiset_of xs) = multiset_of (filter f xs)" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2056 | by (simp add: multiset_of_filter) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2057 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2058 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2059 | "multiset_of xs - multiset_of ys = multiset_of (fold remove1 ys xs)" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2060 | by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2061 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2062 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2063 | "multiset_of xs #\<inter> multiset_of ys = | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2064 | multiset_of (snd (fold (\<lambda>x (ys, zs). | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2065 | if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, [])))" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2066 | proof - | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2067 | have "\<And>zs. multiset_of (snd (fold (\<lambda>x (ys, zs). | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2068 | if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) = | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2069 | (multiset_of xs #\<inter> multiset_of ys) + multiset_of zs" | 
| 51623 | 2070 | by (induct xs arbitrary: ys) | 
| 2071 | (auto simp add: mem_set_multiset_eq inter_add_right1 inter_add_right2 ac_simps) | |
| 2072 | then show ?thesis by simp | |
| 2073 | qed | |
| 2074 | ||
| 2075 | lemma [code]: | |
| 2076 | "multiset_of xs #\<union> multiset_of ys = | |
| 2077 | multiset_of (split append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" | |
| 2078 | proof - | |
| 2079 | have "\<And>zs. multiset_of (split append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) = | |
| 2080 | (multiset_of xs #\<union> multiset_of ys) + multiset_of zs" | |
| 2081 | by (induct xs arbitrary: ys) (simp_all add: multiset_eq_iff) | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2082 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2083 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2084 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2085 | lemma [code_unfold]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2086 | "x \<in># multiset_of xs \<longleftrightarrow> x \<in> set xs" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2087 | by (simp add: in_multiset_of) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2088 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2089 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2090 | "count (multiset_of xs) x = fold (\<lambda>y. if x = y then Suc else id) xs 0" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2091 | proof - | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2092 | have "\<And>n. fold (\<lambda>y. if x = y then Suc else id) xs n = count (multiset_of xs) x + n" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2093 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2094 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2095 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2096 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2097 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2098 | "set_of (multiset_of xs) = set xs" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2099 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2100 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2101 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2102 | "sorted_list_of_multiset (multiset_of xs) = sort xs" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2103 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2104 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2105 | lemma [code]: -- {* not very efficient, but representation-ignorant! *}
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2106 | "multiset_of_set A = multiset_of (sorted_list_of_set A)" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2107 | apply (cases "finite A") | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2108 | apply simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2109 | apply (induct A rule: finite_induct) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2110 | apply (simp_all add: union_commute) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2111 | done | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2112 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2113 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2114 | "mcard (multiset_of xs) = length xs" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2115 | by (simp add: mcard_multiset_of) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2116 | |
| 58425 | 2117 | fun ms_lesseq_impl :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool option" where | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2118 | "ms_lesseq_impl [] ys = Some (ys \<noteq> [])" | 
| 58425 | 2119 | | "ms_lesseq_impl (Cons x xs) ys = (case List.extract (op = x) ys of | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2120 | None \<Rightarrow> None | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2121 | | Some (ys1,_,ys2) \<Rightarrow> ms_lesseq_impl xs (ys1 @ ys2))" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2122 | |
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2123 | lemma ms_lesseq_impl: "(ms_lesseq_impl xs ys = None \<longleftrightarrow> \<not> multiset_of xs \<le> multiset_of ys) \<and> | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2124 | (ms_lesseq_impl xs ys = Some True \<longleftrightarrow> multiset_of xs < multiset_of ys) \<and> | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2125 | (ms_lesseq_impl xs ys = Some False \<longrightarrow> multiset_of xs = multiset_of ys)" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2126 | proof (induct xs arbitrary: ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2127 | case (Nil ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2128 | show ?case by (auto simp: mset_less_empty_nonempty) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2129 | next | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2130 | case (Cons x xs ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2131 | show ?case | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2132 | proof (cases "List.extract (op = x) ys") | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2133 | case None | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2134 | hence x: "x \<notin> set ys" by (simp add: extract_None_iff) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2135 |     {
 | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2136 | assume "multiset_of (x # xs) \<le> multiset_of ys" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2137 | from set_of_mono[OF this] x have False by simp | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2138 | } note nle = this | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2139 | moreover | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2140 |     {
 | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2141 | assume "multiset_of (x # xs) < multiset_of ys" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2142 | hence "multiset_of (x # xs) \<le> multiset_of ys" by auto | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2143 | from nle[OF this] have False . | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2144 | } | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2145 | ultimately show ?thesis using None by auto | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2146 | next | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2147 | case (Some res) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2148 | obtain ys1 y ys2 where res: "res = (ys1,y,ys2)" by (cases res, auto) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2149 | note Some = Some[unfolded res] | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2150 | from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp | 
| 58425 | 2151 |     hence id: "multiset_of ys = multiset_of (ys1 @ ys2) + {#x#}"
 | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2152 | by (auto simp: ac_simps) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2153 | show ?thesis unfolding ms_lesseq_impl.simps | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2154 | unfolding Some option.simps split | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2155 | unfolding id | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2156 | using Cons[of "ys1 @ ys2"] | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2157 | unfolding mset_le_def mset_less_def by auto | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2158 | qed | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2159 | qed | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2160 | |
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2161 | lemma [code]: "multiset_of xs \<le> multiset_of ys \<longleftrightarrow> ms_lesseq_impl xs ys \<noteq> None" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2162 | using ms_lesseq_impl[of xs ys] by (cases "ms_lesseq_impl xs ys", auto) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2163 | |
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2164 | lemma [code]: "multiset_of xs < multiset_of ys \<longleftrightarrow> ms_lesseq_impl xs ys = Some True" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2165 | using ms_lesseq_impl[of xs ys] by (cases "ms_lesseq_impl xs ys", auto) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2166 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2167 | instantiation multiset :: (equal) equal | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2168 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2169 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2170 | definition | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2171 | [code del]: "HOL.equal A (B :: 'a multiset) \<longleftrightarrow> A = B" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2172 | lemma [code]: "HOL.equal (multiset_of xs) (multiset_of ys) \<longleftrightarrow> ms_lesseq_impl xs ys = Some False" | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2173 | unfolding equal_multiset_def | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2174 | using ms_lesseq_impl[of xs ys] by (cases "ms_lesseq_impl xs ys", auto) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2175 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2176 | instance | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2177 | by default (simp add: equal_multiset_def) | 
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
37107diff
changeset | 2178 | end | 
| 49388 | 2179 | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2180 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2181 | "msetsum (multiset_of xs) = listsum xs" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2182 | by (induct xs) (simp_all add: add.commute) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2183 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2184 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2185 | "msetprod (multiset_of xs) = fold times xs 1" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2186 | proof - | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2187 | have "\<And>x. fold times xs x = msetprod (multiset_of xs) * x" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2188 | by (induct xs) (simp_all add: mult.assoc) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2189 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2190 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2191 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2192 | lemma [code]: | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2193 | "size = mcard" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2194 | by (fact size_eq_mcard) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2195 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2196 | text {*
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2197 |   Exercise for the casual reader: add implementations for @{const le_multiset}
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2198 |   and @{const less_multiset} (multiset order).
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2199 | *} | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2200 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2201 | text {* Quickcheck generators *}
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2202 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2203 | definition (in term_syntax) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2204 | msetify :: "'a\<Colon>typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2205 | \<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2206 |   [code_unfold]: "msetify xs = Code_Evaluation.valtermify multiset_of {\<cdot>} xs"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2207 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2208 | notation fcomp (infixl "\<circ>>" 60) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2209 | notation scomp (infixl "\<circ>\<rightarrow>" 60) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2210 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2211 | instantiation multiset :: (random) random | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2212 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2213 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2214 | definition | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2215 | "Quickcheck_Random.random i = Quickcheck_Random.random i \<circ>\<rightarrow> (\<lambda>xs. Pair (msetify xs))" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2216 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2217 | instance .. | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2218 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2219 | end | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2220 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2221 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2222 | no_notation scomp (infixl "\<circ>\<rightarrow>" 60) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2223 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2224 | instantiation multiset :: (full_exhaustive) full_exhaustive | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2225 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2226 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2227 | definition full_exhaustive_multiset :: "('a multiset \<times> (unit \<Rightarrow> term) \<Rightarrow> (bool \<times> term list) option) \<Rightarrow> natural \<Rightarrow> (bool \<times> term list) option"
 | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2228 | where | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2229 | "full_exhaustive_multiset f i = Quickcheck_Exhaustive.full_exhaustive (\<lambda>xs. f (msetify xs)) i" | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2230 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2231 | instance .. | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2232 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2233 | end | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2234 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2235 | hide_const (open) msetify | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2236 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2237 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2238 | subsection {* BNF setup *}
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2239 | |
| 57966 | 2240 | definition rel_mset where | 
| 2241 | "rel_mset R X Y \<longleftrightarrow> (\<exists>xs ys. multiset_of xs = X \<and> multiset_of ys = Y \<and> list_all2 R xs ys)" | |
| 2242 | ||
| 2243 | lemma multiset_of_zip_take_Cons_drop_twice: | |
| 2244 | assumes "length xs = length ys" "j \<le> length xs" | |
| 2245 | shows "multiset_of (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) = | |
| 2246 |     multiset_of (zip xs ys) + {#(x, y)#}"
 | |
| 2247 | using assms | |
| 2248 | proof (induct xs ys arbitrary: x y j rule: list_induct2) | |
| 2249 | case Nil | |
| 2250 | thus ?case | |
| 2251 | by simp | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2252 | next | 
| 57966 | 2253 | case (Cons x xs y ys) | 
| 2254 | thus ?case | |
| 2255 | proof (cases "j = 0") | |
| 2256 | case True | |
| 2257 | thus ?thesis | |
| 2258 | by simp | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2259 | next | 
| 57966 | 2260 | case False | 
| 2261 | then obtain k where k: "j = Suc k" | |
| 2262 | by (case_tac j) simp | |
| 2263 | hence "k \<le> length xs" | |
| 2264 | using Cons.prems by auto | |
| 2265 | hence "multiset_of (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) = | |
| 2266 |       multiset_of (zip xs ys) + {#(x, y)#}"
 | |
| 2267 | by (rule Cons.hyps(2)) | |
| 2268 | thus ?thesis | |
| 2269 | unfolding k by (auto simp: add.commute union_lcomm) | |
| 58425 | 2270 | qed | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2271 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2272 | |
| 57966 | 2273 | lemma ex_multiset_of_zip_left: | 
| 2274 | assumes "length xs = length ys" "multiset_of xs' = multiset_of xs" | |
| 2275 | shows "\<exists>ys'. length ys' = length xs' \<and> multiset_of (zip xs' ys') = multiset_of (zip xs ys)" | |
| 58425 | 2276 | using assms | 
| 57966 | 2277 | proof (induct xs ys arbitrary: xs' rule: list_induct2) | 
| 2278 | case Nil | |
| 2279 | thus ?case | |
| 2280 | by auto | |
| 2281 | next | |
| 2282 | case (Cons x xs y ys xs') | |
| 2283 | obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" | |
| 58425 | 2284 | by (metis Cons.prems in_set_conv_nth list.set_intros(1) multiset_of_eq_setD) | 
| 2285 | ||
| 2286 | def xsa \<equiv> "take j xs' @ drop (Suc j) xs'" | |
| 57966 | 2287 |   have "multiset_of xs' = {#x#} + multiset_of xsa"
 | 
| 2288 | unfolding xsa_def using j_len nth_j | |
| 58247 
98d0f85d247f
enamed drop_Suc_conv_tl and nth_drop' to Cons_nth_drop_Suc
 nipkow parents: 
58098diff
changeset | 2289 | by (metis (no_types) ab_semigroup_add_class.add_ac(1) append_take_drop_id Cons_nth_drop_Suc | 
| 57966 | 2290 | multiset_of.simps(2) union_code union_commute) | 
| 2291 | hence ms_x: "multiset_of xsa = multiset_of xs" | |
| 2292 | by (metis Cons.prems add.commute add_right_imp_eq multiset_of.simps(2)) | |
| 2293 | then obtain ysa where | |
| 2294 | len_a: "length ysa = length xsa" and ms_a: "multiset_of (zip xsa ysa) = multiset_of (zip xs ys)" | |
| 2295 | using Cons.hyps(2) by blast | |
| 2296 | ||
| 2297 | def ys' \<equiv> "take j ysa @ y # drop j ysa" | |
| 2298 | have xs': "xs' = take j xsa @ x # drop j xsa" | |
| 2299 | using ms_x j_len nth_j Cons.prems xsa_def | |
| 58247 
98d0f85d247f
enamed drop_Suc_conv_tl and nth_drop' to Cons_nth_drop_Suc
 nipkow parents: 
58098diff
changeset | 2300 | by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons | 
| 57966 | 2301 | length_drop mcard_multiset_of) | 
| 2302 | have j_len': "j \<le> length xsa" | |
| 2303 | using j_len xs' xsa_def | |
| 2304 | by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) | |
| 2305 | have "length ys' = length xs'" | |
| 2306 | unfolding ys'_def using Cons.prems len_a ms_x | |
| 2307 | by (metis add_Suc_right append_take_drop_id length_Cons length_append multiset_of_eq_length) | |
| 2308 | moreover have "multiset_of (zip xs' ys') = multiset_of (zip (x # xs) (y # ys))" | |
| 2309 | unfolding xs' ys'_def | |
| 2310 | by (rule trans[OF multiset_of_zip_take_Cons_drop_twice]) | |
| 2311 | (auto simp: len_a ms_a j_len' add.commute) | |
| 2312 | ultimately show ?case | |
| 2313 | by blast | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2314 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2315 | |
| 57966 | 2316 | lemma list_all2_reorder_left_invariance: | 
| 2317 | assumes rel: "list_all2 R xs ys" and ms_x: "multiset_of xs' = multiset_of xs" | |
| 2318 | shows "\<exists>ys'. list_all2 R xs' ys' \<and> multiset_of ys' = multiset_of ys" | |
| 2319 | proof - | |
| 2320 | have len: "length xs = length ys" | |
| 2321 | using rel list_all2_conv_all_nth by auto | |
| 2322 | obtain ys' where | |
| 2323 | len': "length xs' = length ys'" and ms_xy: "multiset_of (zip xs' ys') = multiset_of (zip xs ys)" | |
| 2324 | using len ms_x by (metis ex_multiset_of_zip_left) | |
| 2325 | have "list_all2 R xs' ys'" | |
| 2326 | using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: multiset_of_eq_setD) | |
| 2327 | moreover have "multiset_of ys' = multiset_of ys" | |
| 2328 | using len len' ms_xy map_snd_zip multiset_of_map by metis | |
| 2329 | ultimately show ?thesis | |
| 2330 | by blast | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2331 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2332 | |
| 57966 | 2333 | lemma ex_multiset_of: "\<exists>xs. multiset_of xs = X" | 
| 2334 | by (induct X) (simp, metis multiset_of.simps(2)) | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2335 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2336 | bnf "'a multiset" | 
| 57966 | 2337 | map: image_mset | 
| 58425 | 2338 | sets: set_of | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2339 | bd: natLeq | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2340 |   wits: "{#}"
 | 
| 57966 | 2341 | rel: rel_mset | 
| 2342 | proof - | |
| 2343 | show "image_mset id = id" | |
| 2344 | by (rule image_mset.id) | |
| 2345 | next | |
| 2346 | show "\<And>f g. image_mset (g \<circ> f) = image_mset g \<circ> image_mset f" | |
| 2347 | unfolding comp_def by (rule ext) (simp add: image_mset.compositionality comp_def) | |
| 2348 | next | |
| 2349 | fix X :: "'a multiset" | |
| 2350 | show "\<And>f g. (\<And>z. z \<in> set_of X \<Longrightarrow> f z = g z) \<Longrightarrow> image_mset f X = image_mset g X" | |
| 2351 | by (induct X, (simp (no_asm))+, | |
| 2352 | metis One_nat_def Un_iff count_single mem_set_of_iff set_of_union zero_less_Suc) | |
| 2353 | next | |
| 2354 | show "\<And>f. set_of \<circ> image_mset f = op ` f \<circ> set_of" | |
| 2355 | by auto | |
| 2356 | next | |
| 2357 | show "card_order natLeq" | |
| 2358 | by (rule natLeq_card_order) | |
| 2359 | next | |
| 2360 | show "BNF_Cardinal_Arithmetic.cinfinite natLeq" | |
| 2361 | by (rule natLeq_cinfinite) | |
| 2362 | next | |
| 2363 | show "\<And>X. ordLeq3 (card_of (set_of X)) natLeq" | |
| 2364 | by transfer | |
| 2365 | (auto intro!: ordLess_imp_ordLeq simp: finite_iff_ordLess_natLeq[symmetric] multiset_def) | |
| 2366 | next | |
| 2367 | show "\<And>R S. rel_mset R OO rel_mset S \<le> rel_mset (R OO S)" | |
| 2368 | unfolding rel_mset_def[abs_def] OO_def | |
| 2369 | apply clarify | |
| 2370 | apply (rename_tac X Z Y xs ys' ys zs) | |
| 2371 | apply (drule_tac xs = ys' and ys = zs and xs' = ys in list_all2_reorder_left_invariance) | |
| 2372 | by (auto intro: list_all2_trans) | |
| 2373 | next | |
| 2374 | show "\<And>R. rel_mset R = | |
| 2375 |     (BNF_Def.Grp {x. set_of x \<subseteq> {(x, y). R x y}} (image_mset fst))\<inverse>\<inverse> OO
 | |
| 2376 |     BNF_Def.Grp {x. set_of x \<subseteq> {(x, y). R x y}} (image_mset snd)"
 | |
| 2377 | unfolding rel_mset_def[abs_def] BNF_Def.Grp_def OO_def | |
| 2378 | apply (rule ext)+ | |
| 2379 | apply auto | |
| 2380 | apply (rule_tac x = "multiset_of (zip xs ys)" in exI) | |
| 2381 | apply auto[1] | |
| 2382 | apply (metis list_all2_lengthD map_fst_zip multiset_of_map) | |
| 2383 | apply (auto simp: list_all2_iff)[1] | |
| 2384 | apply (metis list_all2_lengthD map_snd_zip multiset_of_map) | |
| 2385 | apply (auto simp: list_all2_iff)[1] | |
| 2386 | apply (rename_tac XY) | |
| 2387 | apply (cut_tac X = XY in ex_multiset_of) | |
| 2388 | apply (erule exE) | |
| 2389 | apply (rename_tac xys) | |
| 2390 | apply (rule_tac x = "map fst xys" in exI) | |
| 2391 | apply (auto simp: multiset_of_map) | |
| 2392 | apply (rule_tac x = "map snd xys" in exI) | |
| 2393 | by (auto simp: multiset_of_map list_all2I subset_eq zip_map_fst_snd) | |
| 2394 | next | |
| 2395 |   show "\<And>z. z \<in> set_of {#} \<Longrightarrow> False"
 | |
| 2396 | by auto | |
| 2397 | qed | |
| 2398 | ||
| 2399 | inductive rel_mset' where | |
| 2400 |   Zero[intro]: "rel_mset' R {#} {#}"
 | |
| 2401 | | Plus[intro]: "\<lbrakk>R a b; rel_mset' R M N\<rbrakk> \<Longrightarrow> rel_mset' R (M + {#a#}) (N + {#b#})"
 | |
| 2402 | ||
| 2403 | lemma rel_mset_Zero: "rel_mset R {#} {#}"
 | |
| 2404 | unfolding rel_mset_def Grp_def by auto | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2405 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2406 | declare multiset.count[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2407 | declare Abs_multiset_inverse[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2408 | declare multiset.count_inverse[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2409 | declare union_preserves_multiset[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2410 | |
| 57966 | 2411 | lemma rel_mset_Plus: | 
| 2412 | assumes ab: "R a b" and MN: "rel_mset R M N" | |
| 2413 | shows "rel_mset R (M + {#a#}) (N + {#b#})"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2414 | proof- | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2415 |   {fix y assume "R a b" and "set_of y \<subseteq> {(x, y). R x y}"
 | 
| 57966 | 2416 |    hence "\<exists>ya. image_mset fst y + {#a#} = image_mset fst ya \<and>
 | 
| 2417 |                image_mset snd y + {#b#} = image_mset snd ya \<and>
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2418 |                set_of ya \<subseteq> {(x, y). R x y}"
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2419 |    apply(intro exI[of _ "y + {#(a,b)#}"]) by auto
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2420 | } | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2421 | thus ?thesis | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2422 | using assms | 
| 57966 | 2423 | unfolding multiset.rel_compp_Grp Grp_def by blast | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2424 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2425 | |
| 57966 | 2426 | lemma rel_mset'_imp_rel_mset: | 
| 2427 | "rel_mset' R M N \<Longrightarrow> rel_mset R M N" | |
| 2428 | apply(induct rule: rel_mset'.induct) | |
| 2429 | using rel_mset_Zero rel_mset_Plus by auto | |
| 2430 | ||
| 2431 | lemma mcard_image_mset[simp]: "mcard (image_mset f M) = mcard M" | |
| 2432 | unfolding size_eq_mcard[symmetric] by (rule size_image_mset) | |
| 2433 | ||
| 2434 | lemma rel_mset_mcard: | |
| 2435 | assumes "rel_mset R M N" | |
| 2436 | shows "mcard M = mcard N" | |
| 2437 | using assms unfolding multiset.rel_compp_Grp Grp_def by auto | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2438 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2439 | lemma multiset_induct2[case_names empty addL addR]: | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2440 | assumes empty: "P {#} {#}"
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2441 | and addL: "\<And>M N a. P M N \<Longrightarrow> P (M + {#a#}) N"
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2442 | and addR: "\<And>M N a. P M N \<Longrightarrow> P M (N + {#a#})"
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2443 | shows "P M N" | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2444 | apply(induct N rule: multiset_induct) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2445 | apply(induct M rule: multiset_induct, rule empty, erule addL) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2446 | apply(induct M rule: multiset_induct, erule addR, erule addR) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2447 | done | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2448 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2449 | lemma multiset_induct2_mcard[consumes 1, case_names empty add]: | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2450 | assumes c: "mcard M = mcard N" | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2451 | and empty: "P {#} {#}"
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2452 | and add: "\<And>M N a b. P M N \<Longrightarrow> P (M + {#a#}) (N + {#b#})"
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2453 | shows "P M N" | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2454 | using c proof(induct M arbitrary: N rule: measure_induct_rule[of mcard]) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2455 | case (less M) show ?case | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2456 |   proof(cases "M = {#}")
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2457 |     case True hence "N = {#}" using less.prems by auto
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2458 | thus ?thesis using True empty by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2459 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2460 |     case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2461 |     have "N \<noteq> {#}" using False less.prems by auto
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2462 |     then obtain N1 b where N: "N = N1 + {#b#}" by (metis multi_nonempty_split)
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2463 | have "mcard M1 = mcard N1" using less.prems unfolding M N by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2464 | thus ?thesis using M N less.hyps add by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2465 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2466 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2467 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2468 | lemma msed_map_invL: | 
| 57966 | 2469 | assumes "image_mset f (M + {#a#}) = N"
 | 
| 2470 | shows "\<exists>N1. N = N1 + {#f a#} \<and> image_mset f M = N1"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2471 | proof- | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2472 | have "f a \<in># N" | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2473 |   using assms multiset.set_map[of f "M + {#a#}"] by auto
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2474 |   then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
 | 
| 57966 | 2475 | have "image_mset f M = N1" using assms unfolding N by simp | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2476 | thus ?thesis using N by blast | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2477 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2478 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2479 | lemma msed_map_invR: | 
| 57966 | 2480 | assumes "image_mset f M = N + {#b#}"
 | 
| 2481 | shows "\<exists>M1 a. M = M1 + {#a#} \<and> f a = b \<and> image_mset f M1 = N"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2482 | proof- | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2483 | obtain a where a: "a \<in># M" and fa: "f a = b" | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2484 | using multiset.set_map[of f M] unfolding assms | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2485 | by (metis image_iff mem_set_of_iff union_single_eq_member) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2486 |   then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
 | 
| 57966 | 2487 | have "image_mset f M1 = N" using assms unfolding M fa[symmetric] by simp | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2488 | thus ?thesis using M fa by blast | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2489 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2490 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2491 | lemma msed_rel_invL: | 
| 57966 | 2492 | assumes "rel_mset R (M + {#a#}) N"
 | 
| 2493 | shows "\<exists>N1 b. N = N1 + {#b#} \<and> R a b \<and> rel_mset R M N1"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2494 | proof- | 
| 57966 | 2495 |   obtain K where KM: "image_mset fst K = M + {#a#}"
 | 
| 2496 |   and KN: "image_mset snd K = N" and sK: "set_of K \<subseteq> {(a, b). R a b}"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2497 | using assms | 
| 57966 | 2498 | unfolding multiset.rel_compp_Grp Grp_def by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2499 |   obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
 | 
| 57966 | 2500 | and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto | 
| 2501 |   obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "image_mset snd K1 = N1"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2502 | using msed_map_invL[OF KN[unfolded K]] by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2503 | have Rab: "R a (snd ab)" using sK a unfolding K by auto | 
| 57966 | 2504 | have "rel_mset R M N1" using sK K1M K1N1 | 
| 2505 | unfolding K multiset.rel_compp_Grp Grp_def by auto | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2506 | thus ?thesis using N Rab by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2507 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2508 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2509 | lemma msed_rel_invR: | 
| 57966 | 2510 | assumes "rel_mset R M (N + {#b#})"
 | 
| 2511 | shows "\<exists>M1 a. M = M1 + {#a#} \<and> R a b \<and> rel_mset R M1 N"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2512 | proof- | 
| 57966 | 2513 |   obtain K where KN: "image_mset snd K = N + {#b#}"
 | 
| 2514 |   and KM: "image_mset fst K = M" and sK: "set_of K \<subseteq> {(a, b). R a b}"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2515 | using assms | 
| 57966 | 2516 | unfolding multiset.rel_compp_Grp Grp_def by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2517 |   obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
 | 
| 57966 | 2518 | and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto | 
| 2519 |   obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "image_mset fst K1 = M1"
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2520 | using msed_map_invL[OF KM[unfolded K]] by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2521 | have Rab: "R (fst ab) b" using sK b unfolding K by auto | 
| 57966 | 2522 | have "rel_mset R M1 N" using sK K1N K1M1 | 
| 2523 | unfolding K multiset.rel_compp_Grp Grp_def by auto | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2524 | thus ?thesis using M Rab by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2525 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2526 | |
| 57966 | 2527 | lemma rel_mset_imp_rel_mset': | 
| 2528 | assumes "rel_mset R M N" | |
| 2529 | shows "rel_mset' R M N" | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2530 | using assms proof(induct M arbitrary: N rule: measure_induct_rule[of mcard]) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2531 | case (less M) | 
| 57966 | 2532 | have c: "mcard M = mcard N" using rel_mset_mcard[OF less.prems] . | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2533 | show ?case | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2534 |   proof(cases "M = {#}")
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2535 |     case True hence "N = {#}" using c by simp
 | 
| 57966 | 2536 | thus ?thesis using True rel_mset'.Zero by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2537 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2538 |     case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
 | 
| 57966 | 2539 |     obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "rel_mset R M1 N1"
 | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2540 | using msed_rel_invL[OF less.prems[unfolded M]] by auto | 
| 57966 | 2541 | have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp | 
| 2542 | thus ?thesis using rel_mset'.Plus[of R a b, OF R] unfolding M N by simp | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2543 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2544 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2545 | |
| 57966 | 2546 | lemma rel_mset_rel_mset': | 
| 2547 | "rel_mset R M N = rel_mset' R M N" | |
| 2548 | using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto | |
| 2549 | ||
| 2550 | (* The main end product for rel_mset: inductive characterization *) | |
| 2551 | theorems rel_mset_induct[case_names empty add, induct pred: rel_mset] = | |
| 2552 | rel_mset'.induct[unfolded rel_mset_rel_mset'[symmetric]] | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2553 | |
| 56656 | 2554 | |
| 2555 | subsection {* Size setup *}
 | |
| 2556 | ||
| 57966 | 2557 | lemma multiset_size_o_map: "size_multiset g \<circ> image_mset f = size_multiset (g \<circ> f)" | 
| 2558 | unfolding o_apply by (rule ext) (induct_tac, auto) | |
| 56656 | 2559 | |
| 2560 | setup {*
 | |
| 2561 | BNF_LFP_Size.register_size_global @{type_name multiset} @{const_name size_multiset}
 | |
| 2562 |   @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single
 | |
| 2563 | size_union} | |
| 2564 |   @{thms multiset_size_o_map}
 | |
| 2565 | *} | |
| 2566 | ||
| 2567 | hide_const (open) wcount | |
| 2568 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2569 | end |