| author | wenzelm | 
| Tue, 08 Mar 2016 18:38:29 +0100 | |
| changeset 62560 | 498f6ff16804 | 
| parent 62430 | 9527ff088c15 | 
| child 62537 | 7a9aa69f9b38 | 
| 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 | 
| 59813 | 4 | Author: Jasmin Blanchette, Inria, LORIA, MPII | 
| 5 | Author: Dmitriy Traytel, TU Muenchen | |
| 6 | Author: Mathias Fleury, MPII | |
| 10249 | 7 | *) | 
| 8 | ||
| 60500 | 9 | section \<open>(Finite) multisets\<close> | 
| 10249 | 10 | |
| 15131 | 11 | theory Multiset | 
| 51599 
1559e9266280
optionalized very specific code setup for multisets
 haftmann parents: 
51548diff
changeset | 12 | imports Main | 
| 15131 | 13 | begin | 
| 10249 | 14 | |
| 60500 | 15 | subsection \<open>The type of multisets\<close> | 
| 10249 | 16 | |
| 60606 | 17 | definition "multiset = {f :: 'a \<Rightarrow> nat. finite {x. f x > 0}}"
 | 
| 18 | ||
| 19 | typedef 'a multiset = "multiset :: ('a \<Rightarrow> 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 | 20 | morphisms count Abs_multiset | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45608diff
changeset | 21 | unfolding multiset_def | 
| 10249 | 22 | proof | 
| 45694 
4a8743618257
prefer typedef without extra definition and alternative name;
 wenzelm parents: 
45608diff
changeset | 23 |   show "(\<lambda>x. 0::nat) \<in> {f. finite {x. f x > 0}}" by simp
 | 
| 10249 | 24 | qed | 
| 25 | ||
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 26 | setup_lifting type_definition_multiset | 
| 19086 | 27 | |
| 60606 | 28 | lemma multiset_eq_iff: "M = N \<longleftrightarrow> (\<forall>a. count M a = count N a)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 29 | by (simp only: count_inject [symmetric] fun_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 30 | |
| 60606 | 31 | lemma multiset_eqI: "(\<And>x. count A x = count B x) \<Longrightarrow> A = B" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 32 | 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 | 33 | |
| 60606 | 34 | text \<open>Preservation of the representing set @{term multiset}.\<close>
 | 
| 35 | ||
| 36 | lemma const0_in_multiset: "(\<lambda>a. 0) \<in> 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 | 37 | 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 | 38 | |
| 60606 | 39 | lemma only1_in_multiset: "(\<lambda>b. if b = a then n else 0) \<in> 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 | 40 | 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 | 41 | |
| 60606 | 42 | lemma union_preserves_multiset: "M \<in> multiset \<Longrightarrow> N \<in> multiset \<Longrightarrow> (\<lambda>a. M a + N a) \<in> 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 | 43 | 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 | 44 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 46 | 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 | 47 | 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 | 48 | 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 | 49 |   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 | 50 | 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 | 51 | 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 | 52 | 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 | 53 | 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 | 54 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 55 | 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 | 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>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 | 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. (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 | 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 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 65 | 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 | 66 | 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 | 67 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | |
| 60500 | 69 | subsection \<open>Representing multisets\<close> | 
| 70 | ||
| 71 | text \<open>Multiset enumeration\<close> | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 72 | |
| 48008 | 73 | instantiation multiset :: (type) cancel_comm_monoid_add | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 74 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 75 | |
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 76 | lift_definition zero_multiset :: "'a multiset" is "\<lambda>a. 0" | 
| 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 77 | by (rule const0_in_multiset) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 78 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 80 | "Mempty \<equiv> 0" | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 81 | |
| 60606 | 82 | lift_definition plus_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" is "\<lambda>M N. (\<lambda>a. M a + N a)" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 83 | by (rule union_preserves_multiset) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 84 | |
| 60606 | 85 | lift_definition minus_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" is "\<lambda> M N. \<lambda>a. M a - N a" | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59813diff
changeset | 86 | by (rule diff_preserves_multiset) | 
| 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59813diff
changeset | 87 | |
| 48008 | 88 | instance | 
| 60678 | 89 | by (standard; transfer; simp add: fun_eq_iff) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 90 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 91 | end | 
| 10249 | 92 | |
| 60606 | 93 | lift_definition single :: "'a \<Rightarrow> 'a multiset" is "\<lambda>a b. if b = a then 1 else 0" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 94 | by (rule only1_in_multiset) | 
| 15869 | 95 | |
| 26145 | 96 | syntax | 
| 60606 | 97 |   "_multiset" :: "args \<Rightarrow> 'a multiset"    ("{#(_)#}")
 | 
| 25507 | 98 | translations | 
| 99 |   "{#x, xs#}" == "{#x#} + {#xs#}"
 | |
| 100 |   "{#x#}" == "CONST single x"
 | |
| 101 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 102 | lemma count_empty [simp]: "count {#} a = 0"
 | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 103 | by (simp add: zero_multiset.rep_eq) | 
| 10249 | 104 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 105 | 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 | 106 | by (simp add: single.rep_eq) | 
| 29901 | 107 | |
| 10249 | 108 | |
| 60500 | 109 | subsection \<open>Basic operations\<close> | 
| 110 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 111 | subsubsection \<open>Conversion to set and membership\<close> | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 112 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 113 | definition set_mset :: "'a multiset \<Rightarrow> 'a set" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 114 |   where "set_mset M = {x. count M x > 0}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 115 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 116 | abbreviation Melem :: "'a \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<in>#" 50) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 117 | where "a \<in># M \<equiv> a \<in> set_mset M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 118 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 119 | abbreviation not_Melem :: "'a \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<notin>#" 50) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 120 | where "a \<notin># M \<equiv> a \<notin> set_mset M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 121 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 122 | context | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 123 | begin | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 124 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 125 | qualified abbreviation Ball :: "'a multiset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 126 | where "Ball M \<equiv> Set.Ball (set_mset M)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 127 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 128 | qualified abbreviation Bex :: "'a multiset \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 129 | where "Bex M \<equiv> Set.Bex (set_mset M)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 130 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 131 | end | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 132 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 133 | syntax | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 134 |   "_MBall"       :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<forall>_\<in>#_./ _)" [0, 0, 10] 10)
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 135 |   "_MBex"        :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<exists>_\<in>#_./ _)" [0, 0, 10] 10)
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 136 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 137 | translations | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 138 | "\<forall>x\<in>#A. P" \<rightleftharpoons> "CONST Multiset.Ball A (\<lambda>x. P)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 139 | "\<exists>x\<in>#A. P" \<rightleftharpoons> "CONST Multiset.Bex A (\<lambda>x. P)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 140 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 141 | lemma count_eq_zero_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 142 | "count M x = 0 \<longleftrightarrow> x \<notin># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 143 | by (auto simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 144 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 145 | lemma not_in_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 146 | "x \<notin># M \<longleftrightarrow> count M x = 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 147 | by (auto simp add: count_eq_zero_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 148 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 149 | lemma count_greater_zero_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 150 | "count M x > 0 \<longleftrightarrow> x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 151 | by (auto simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 152 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 153 | lemma count_inI: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 154 | assumes "count M x = 0 \<Longrightarrow> False" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 155 | shows "x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 156 | proof (rule ccontr) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 157 | assume "x \<notin># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 158 | with assms show False by (simp add: not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 159 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 160 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 161 | lemma in_countE: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 162 | assumes "x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 163 | obtains n where "count M x = Suc n" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 164 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 165 | from assms have "count M x > 0" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 166 | then obtain n where "count M x = Suc n" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 167 | using gr0_conv_Suc by blast | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 168 | with that show thesis . | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 169 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 170 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 171 | lemma count_greater_eq_Suc_zero_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 172 | "count M x \<ge> Suc 0 \<longleftrightarrow> x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 173 | by (simp add: Suc_le_eq) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 174 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 175 | lemma count_greater_eq_one_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 176 | "count M x \<ge> 1 \<longleftrightarrow> x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 177 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 178 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 179 | lemma set_mset_empty [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 180 |   "set_mset {#} = {}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 181 | by (simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 182 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 183 | lemma set_mset_single [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 184 |   "set_mset {#b#} = {b}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 185 | by (simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 186 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 187 | lemma set_mset_eq_empty_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 188 |   "set_mset M = {} \<longleftrightarrow> M = {#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 189 | by (auto simp add: multiset_eq_iff count_eq_zero_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 190 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 191 | lemma finite_set_mset [iff]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 192 | "finite (set_mset M)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 193 | using count [of M] by (simp add: multiset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 194 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 195 | |
| 60500 | 196 | subsubsection \<open>Union\<close> | 
| 10249 | 197 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 198 | lemma count_union [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 199 | "count (M + N) a = count M a + count N a" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 200 | by (simp add: plus_multiset.rep_eq) | 
| 10249 | 201 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 202 | lemma set_mset_union [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 203 | "set_mset (M + N) = set_mset M \<union> set_mset N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 204 | by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_union) simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 205 | |
| 10249 | 206 | |
| 60500 | 207 | subsubsection \<open>Difference\<close> | 
| 10249 | 208 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 209 | instance multiset :: (type) comm_monoid_diff | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 210 | by standard (transfer; simp add: fun_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 211 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 212 | lemma count_diff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 213 | "count (M - N) a = count M a - count N a" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 214 | 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 | 215 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 216 | lemma in_diff_count: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 217 | "a \<in># M - N \<longleftrightarrow> count N a < count M a" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 218 | by (simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 219 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 220 | lemma count_in_diffI: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 221 | assumes "\<And>n. count N x = n + count M x \<Longrightarrow> False" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 222 | shows "x \<in># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 223 | proof (rule ccontr) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 224 | assume "x \<notin># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 225 | then have "count N x = (count N x - count M x) + count M x" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 226 | by (simp add: in_diff_count not_less) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 227 | with assms show False by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 228 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 229 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 230 | lemma in_diff_countE: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 231 | assumes "x \<in># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 232 | obtains n where "count M x = Suc n + count N x" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 233 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 234 | from assms have "count M x - count N x > 0" by (simp add: in_diff_count) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 235 | then have "count M x > count N x" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 236 | then obtain n where "count M x = Suc n + count N x" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 237 | using less_iff_Suc_add by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 238 | with that show thesis . | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 239 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 240 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 241 | lemma in_diffD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 242 | assumes "a \<in># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 243 | shows "a \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 244 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 245 | have "0 \<le> count N a" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 246 | also from assms have "count N a < count M a" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 247 | by (simp add: in_diff_count) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 248 | finally show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 249 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 250 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 251 | lemma set_mset_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 252 |   "set_mset (M - N) = {a. count N a < count M a}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 253 | by (simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 254 | |
| 17161 | 255 | lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
 | 
| 52289 | 256 | by rule (fact Groups.diff_zero, fact Groups.zero_diff) | 
| 36903 | 257 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 258 | lemma diff_cancel [simp]: "A - A = {#}"
 | 
| 52289 | 259 | by (fact Groups.diff_cancel) | 
| 10249 | 260 | |
| 36903 | 261 | lemma diff_union_cancelR [simp]: "M + N - N = (M::'a multiset)" | 
| 52289 | 262 | by (fact add_diff_cancel_right') | 
| 10249 | 263 | |
| 36903 | 264 | lemma diff_union_cancelL [simp]: "N + M - N = (M::'a multiset)" | 
| 52289 | 265 | 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 | 266 | |
| 52289 | 267 | lemma diff_right_commute: | 
| 60606 | 268 | fixes M N Q :: "'a multiset" | 
| 269 | shows "M - N - Q = M - Q - N" | |
| 52289 | 270 | by (fact diff_right_commute) | 
| 271 | ||
| 272 | lemma diff_add: | |
| 60606 | 273 | fixes M N Q :: "'a multiset" | 
| 274 | shows "M - (N + Q) = M - N - Q" | |
| 52289 | 275 | by (rule sym) (fact diff_diff_add) | 
| 58425 | 276 | |
| 60606 | 277 | lemma insert_DiffM: "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 | 278 | 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 | 279 | |
| 60606 | 280 | lemma insert_DiffM2 [simp]: "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 | 281 | 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 | 282 | |
| 60606 | 283 | lemma diff_union_swap: "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 | 284 | 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 | 285 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 286 | lemma diff_union_single_conv: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 287 |   "a \<in># J \<Longrightarrow> I + J - {#a#} = I + (J - {#a#})"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 288 | by (simp add: multiset_eq_iff Suc_le_eq) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 289 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 290 | lemma mset_add [elim?]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 291 | assumes "a \<in># A" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 292 |   obtains B where "A = B + {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 293 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 294 |   from assms have "A = (A - {#a#}) + {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 295 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 296 | with that show thesis . | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 297 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 298 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 299 | lemma union_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 300 | "a \<in># A + B \<longleftrightarrow> a \<in># A \<or> a \<in># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 301 | by auto | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 302 | |
| 10249 | 303 | |
| 60500 | 304 | subsubsection \<open>Equality of multisets\<close> | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 305 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 306 | 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 | 307 | 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 | 308 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 309 | 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 | 310 | 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 | 311 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 312 | lemma 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 | 313 | 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 | 314 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 315 | 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 | 316 | 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 | 317 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 318 | 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 | 319 | 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 | 320 | |
| 60606 | 321 | lemma diff_single_trivial: "\<not> x \<in># M \<Longrightarrow> M - {#x#} = M"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 322 | by (auto simp add: multiset_eq_iff not_in_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 323 | |
| 60606 | 324 | lemma diff_single_eq_union: "x \<in># M \<Longrightarrow> M - {#x#} = N \<longleftrightarrow> M = N + {#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 | 325 | 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 | 326 | |
| 60606 | 327 | lemma union_single_eq_diff: "M + {#x#} = N \<Longrightarrow> M = N - {#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 | 328 | 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 | 329 | |
| 60606 | 330 | lemma union_single_eq_member: "M + {#x#} = N \<Longrightarrow> x \<in># N"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 331 | 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 | 332 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 333 | lemma union_is_single: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 334 |   "M + N = {#a#} \<longleftrightarrow> M = {#a#} \<and> N = {#} \<or> M = {#} \<and> N = {#a#}"
 | 
| 60606 | 335 | (is "?lhs = ?rhs") | 
| 46730 | 336 | proof | 
| 60606 | 337 | show ?lhs if ?rhs using that by auto | 
| 338 | show ?rhs if ?lhs | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 339 | by (metis Multiset.diff_cancel add.commute add_diff_cancel_left' diff_add_zero diff_single_trivial insert_DiffM that) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 340 | 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 | 341 | |
| 60606 | 342 | lemma single_is_union: "{#a#} = M + N \<longleftrightarrow> {#a#} = M \<and> N = {#} \<or> M = {#} \<and> {#a#} = N"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 343 |   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 | 344 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma add_eq_conv_diff: | 
| 60606 | 346 |   "M + {#a#} = N + {#b#} \<longleftrightarrow> M = N \<and> a = b \<or> M = N - {#a#} + {#b#} \<and> N = M - {#b#} + {#a#}"
 | 
| 347 | (is "?lhs \<longleftrightarrow> ?rhs") | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44339diff
changeset | 348 | (* 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 | 349 | proof | 
| 60606 | 350 | show ?lhs if ?rhs | 
| 351 | using that | |
| 352 |     by (auto simp add: add.assoc add.commute [of "{#b#}"])
 | |
| 353 | (drule sym, simp add: add.assoc [symmetric]) | |
| 354 | show ?rhs if ?lhs | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 355 | proof (cases "a = b") | 
| 60500 | 356 | case True with \<open>?lhs\<close> show ?thesis by simp | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 357 | 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 | 358 | case False | 
| 60500 | 359 |     from \<open>?lhs\<close> have "a \<in># N + {#b#}" by (rule union_single_eq_member)
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 360 | with False have "a \<in># N" by auto | 
| 60500 | 361 |     moreover from \<open>?lhs\<close> have "M = N + {#b#} - {#a#}" by (rule union_single_eq_diff)
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 362 | 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 | 363 |     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 | 364 | 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 | 365 | 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 | 366 | |
| 58425 | 367 | 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 | 368 |   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 | 369 | 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 | 370 | 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 | 371 | 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 | 372 |   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 | 373 |   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 | 374 |   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 | 375 | 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 | 376 | 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 | 377 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 378 | 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 | 379 |   "(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 | 380 |     (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 | 381 | 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 | 382 | |
| 60606 | 383 | lemma multi_member_split: "x \<in># M \<Longrightarrow> \<exists>A. M = A + {#x#}"
 | 
| 60678 | 384 |   by (rule exI [where x = "M - {#x#}"]) simp
 | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 385 | |
| 58425 | 386 | lemma multiset_add_sub_el_shuffle: | 
| 60606 | 387 | assumes "c \<in># B" | 
| 388 | and "b \<noteq> c" | |
| 58098 | 389 |   shows "B - {#c#} + {#b#} = B + {#b#} - {#c#}"
 | 
| 390 | proof - | |
| 60500 | 391 |   from \<open>c \<in># B\<close> obtain A where B: "B = A + {#c#}"
 | 
| 58098 | 392 | by (blast dest: multi_member_split) | 
| 393 |   have "A + {#b#} = A + {#b#} + {#c#} - {#c#}" by simp
 | |
| 58425 | 394 |   then have "A + {#b#} = A + {#c#} + {#b#} - {#c#}"
 | 
| 58098 | 395 | by (simp add: ac_simps) | 
| 396 | then show ?thesis using B by simp | |
| 397 | qed | |
| 398 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 399 | |
| 60500 | 400 | subsubsection \<open>Pointwise ordering induced by count\<close> | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 401 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 402 | definition subseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subseteq>#" 50) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 403 | where "A \<subseteq># B = (\<forall>a. count A a \<le> count B a)" | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 404 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 405 | definition subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subset>#" 50) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 406 | where "A \<subset># B = (A \<subseteq># B \<and> A \<noteq> B)" | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 407 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 408 | abbreviation (input) supseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<supseteq>#" 50) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 409 | where "supseteq_mset A B \<equiv> B \<subseteq># A" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 410 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 411 | abbreviation (input) supset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<supset>#" 50) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 412 | where "supset_mset A B \<equiv> B \<subset># A" | 
| 62208 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 413 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 414 | notation (input) | 
| 62208 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 415 | subseteq_mset (infix "\<le>#" 50) and | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 416 | supseteq_mset (infix "\<ge>#" 50) | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 417 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 418 | notation (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 419 | subseteq_mset (infix "<=#" 50) and | 
| 62208 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 420 | subset_mset (infix "<#" 50) and | 
| 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 421 | supseteq_mset (infix ">=#" 50) and | 
| 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 422 | supset_mset (infix ">#" 50) | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 423 | |
| 60606 | 424 | interpretation subset_mset: ordered_ab_semigroup_add_imp_le "op +" "op -" "op \<subseteq>#" "op \<subset>#" | 
| 60678 | 425 | by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 426 | -- \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 427 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 428 | lemma mset_less_eqI: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 429 | "(\<And>a. count A a \<le> count B a) \<Longrightarrow> A \<subseteq># B" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 430 | by (simp add: subseteq_mset_def) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 431 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 432 | lemma mset_less_eq_count: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 433 | "A \<subseteq># B \<Longrightarrow> count A a \<le> count B a" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 434 | by (simp add: subseteq_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 435 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 436 | lemma mset_le_exists_conv: "(A::'a multiset) \<subseteq># B \<longleftrightarrow> (\<exists>C. B = A + C)" | 
| 60678 | 437 | unfolding subseteq_mset_def | 
| 438 | apply (rule iffI) | |
| 439 | apply (rule exI [where x = "B - A"]) | |
| 440 | apply (auto intro: multiset_eq_iff [THEN iffD2]) | |
| 441 | done | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 442 | |
| 62376 
85f38d5f8807
Rename ordered_comm_monoid_add to ordered_cancel_comm_monoid_add. Introduce ordreed_comm_monoid_add, canonically_ordered_comm_monoid and dioid. Setup nat, entat and ennreal as dioids.
 hoelzl parents: 
62366diff
changeset | 443 | interpretation subset_mset: ordered_cancel_comm_monoid_diff "op +" 0 "op \<le>#" "op <#" "op -" | 
| 60678 | 444 | by standard (simp, fact mset_le_exists_conv) | 
| 52289 | 445 | |
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62376diff
changeset | 446 | declare subset_mset.zero_order[simp del] | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62376diff
changeset | 447 | -- \<open>this removes some simp rules not in the usual order for multisets\<close> | 
| 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62376diff
changeset | 448 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 449 | lemma mset_le_mono_add_right_cancel [simp]: "(A::'a multiset) + C \<subseteq># B + C \<longleftrightarrow> A \<subseteq># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 450 | by (fact subset_mset.add_le_cancel_right) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 451 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 452 | lemma mset_le_mono_add_left_cancel [simp]: "C + (A::'a multiset) \<subseteq># C + B \<longleftrightarrow> A \<subseteq># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 453 | by (fact subset_mset.add_le_cancel_left) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 454 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 455 | lemma mset_le_mono_add: "(A::'a multiset) \<subseteq># B \<Longrightarrow> C \<subseteq># D \<Longrightarrow> A + C \<subseteq># B + D" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 456 | by (fact subset_mset.add_mono) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 457 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 458 | lemma mset_le_add_left [simp]: "(A::'a multiset) \<subseteq># A + B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 459 | unfolding subseteq_mset_def by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 460 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 461 | lemma mset_le_add_right [simp]: "B \<subseteq># (A::'a multiset) + B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 462 | unfolding subseteq_mset_def by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 463 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 464 | lemma single_subset_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 465 |   "{#a#} \<subseteq># M \<longleftrightarrow> a \<in># M"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 466 | by (auto simp add: subseteq_mset_def Suc_le_eq) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 467 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 468 | lemma mset_le_single: "a \<in># B \<Longrightarrow> {#a#} \<subseteq># B"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 469 | by (simp add: subseteq_mset_def Suc_le_eq) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 470 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 471 | lemma multiset_diff_union_assoc: | 
| 60606 | 472 | fixes A B C D :: "'a multiset" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 473 | shows "C \<subseteq># B \<Longrightarrow> A + B - C = A + (B - C)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 474 | by (fact subset_mset.diff_add_assoc) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 475 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 476 | lemma mset_le_multiset_union_diff_commute: | 
| 60606 | 477 | fixes A B C D :: "'a multiset" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 478 | shows "B \<subseteq># A \<Longrightarrow> A - B + C = A + C - B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 479 | by (fact subset_mset.add_diff_assoc2) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 480 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 481 | lemma diff_le_self[simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 482 | "(M::'a multiset) - N \<subseteq># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 483 | by (simp add: subseteq_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 484 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 485 | lemma mset_leD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 486 | assumes "A \<subseteq># B" and "x \<in># A" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 487 | shows "x \<in># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 488 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 489 | from \<open>x \<in># A\<close> have "count A x > 0" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 490 | also from \<open>A \<subseteq># B\<close> have "count A x \<le> count B x" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 491 | by (simp add: subseteq_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 492 | finally show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 493 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 494 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 495 | lemma mset_lessD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 496 | "A \<subset># B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 497 | by (auto intro: mset_leD [of A]) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 498 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 499 | lemma set_mset_mono: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 500 | "A \<subseteq># B \<Longrightarrow> set_mset A \<subseteq> set_mset B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 501 | by (metis mset_leD subsetI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 502 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 503 | lemma mset_le_insertD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 504 |   "A + {#x#} \<subseteq># B \<Longrightarrow> x \<in># B \<and> A \<subset># B"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 505 | 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 | 506 | apply (simp add: mset_leD) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 507 | apply (clarsimp simp: subset_mset_def subseteq_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 508 | apply safe | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 509 | apply (erule_tac x = a in allE) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 510 | apply (auto split: if_split_asm) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 511 | 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 | 512 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 513 | lemma mset_less_insertD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 514 |   "A + {#x#} \<subset># B \<Longrightarrow> x \<in># B \<and> A \<subset># B"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 515 | by (rule mset_le_insertD) simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 516 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 517 | lemma mset_less_of_empty[simp]: "A \<subset># {#} \<longleftrightarrow> False"
 | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 518 | by (auto simp add: subseteq_mset_def subset_mset_def multiset_eq_iff) | 
| 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 519 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 520 | lemma empty_le [simp]: "{#} \<subseteq># A"
 | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 521 | unfolding mset_le_exists_conv by auto | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 522 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 523 | lemma insert_subset_eq_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 524 |   "{#a#} + A \<subseteq># B \<longleftrightarrow> a \<in># B \<and> A \<subseteq># B - {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 525 | using le_diff_conv2 [of "Suc 0" "count B a" "count A a"] | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 526 | apply (auto simp add: subseteq_mset_def not_in_iff Suc_le_eq) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 527 | apply (rule ccontr) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 528 | apply (auto simp add: not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 529 | done | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 530 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 531 | lemma insert_union_subset_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 532 |   "{#a#} + A \<subset># B \<longleftrightarrow> a \<in># B \<and> A \<subset># B - {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 533 | by (auto simp add: insert_subset_eq_iff subset_mset_def insert_DiffM) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 534 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 535 | lemma subset_eq_diff_conv: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 536 | "A - C \<subseteq># B \<longleftrightarrow> A \<subseteq># B + C" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 537 | by (simp add: subseteq_mset_def le_diff_conv) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 538 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 539 | lemma le_empty [simp]: "M \<subseteq># {#} \<longleftrightarrow> M = {#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 540 | unfolding mset_le_exists_conv by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 541 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 542 | lemma multi_psub_of_add_self[simp]: "A \<subset># A + {#x#}"
 | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 543 | by (auto simp: subset_mset_def subseteq_mset_def) | 
| 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 544 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 545 | lemma multi_psub_self[simp]: "(A::'a multiset) \<subset># A = False" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 546 | 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 | 547 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 548 | lemma mset_less_add_bothsides: "N + {#x#} \<subset># M + {#x#} \<Longrightarrow> N \<subset># M"
 | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 549 | by (fact subset_mset.add_less_imp_less_right) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 550 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 551 | lemma mset_less_empty_nonempty: "{#} \<subset># S \<longleftrightarrow> S \<noteq> {#}"
 | 
| 62378 
85ed00c1fe7c
generalize more theorems to support enat and ennreal
 hoelzl parents: 
62376diff
changeset | 552 | by (fact subset_mset.zero_less_iff_neq_zero) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 553 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 554 | lemma mset_less_diff_self: "c \<in># B \<Longrightarrow> B - {#c#} \<subset># B"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 555 | by (auto simp: subset_mset_def elim: mset_add) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 556 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 557 | |
| 60500 | 558 | subsubsection \<open>Intersection\<close> | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 559 | |
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 560 | definition inf_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where | 
| 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 561 | multiset_inter_def: "inf_subset_mset A B = A - (A - B)" | 
| 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 562 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 563 | interpretation subset_mset: semilattice_inf inf_subset_mset "op \<subseteq>#" "op \<subset>#" | 
| 46921 | 564 | proof - | 
| 60678 | 565 | have [simp]: "m \<le> n \<Longrightarrow> m \<le> q \<Longrightarrow> m \<le> n - (n - q)" for m n q :: nat | 
| 566 | by arith | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 567 | show "class.semilattice_inf op #\<inter> op \<subseteq># op \<subset>#" | 
| 60678 | 568 | by standard (auto simp add: multiset_inter_def subseteq_mset_def) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 569 | qed | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 570 | -- \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 571 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 572 | lemma multiset_inter_count [simp]: | 
| 60606 | 573 | fixes A B :: "'a multiset" | 
| 574 | shows "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 | 575 | by (simp add: multiset_inter_def) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 576 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 577 | lemma set_mset_inter [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 578 | "set_mset (A #\<inter> B) = set_mset A \<inter> set_mset B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 579 | by (simp only: set_eq_iff count_greater_zero_iff [symmetric] multiset_inter_count) simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 580 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 581 | lemma diff_intersect_left_idem [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 582 | "M - M #\<inter> N = M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 583 | by (simp add: multiset_eq_iff min_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 584 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 585 | lemma diff_intersect_right_idem [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 586 | "M - N #\<inter> M = M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 587 | by (simp add: multiset_eq_iff min_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 588 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 589 | lemma multiset_inter_single: "a \<noteq> b \<Longrightarrow> {#a#} #\<inter> {#b#} = {#}"
 | 
| 46730 | 590 | 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 | 591 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 592 | lemma multiset_union_diff_commute: | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 593 |   assumes "B #\<inter> C = {#}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 594 | 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 | 595 | proof (rule multiset_eqI) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 596 | fix x | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 597 | from assms have "min (count B x) (count C x) = 0" | 
| 46730 | 598 | by (auto simp add: multiset_eq_iff) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 599 | then have "count B x = 0 \<or> count C x = 0" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 600 | unfolding min_def by (auto split: if_splits) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 601 | then show "count (A + B - C) x = count (A - C + B) x" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 602 | by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 603 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 604 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 605 | lemma disjunct_not_in: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 606 |   "A #\<inter> B = {#} \<longleftrightarrow> (\<forall>a. a \<notin># A \<or> a \<notin># B)" (is "?P \<longleftrightarrow> ?Q")
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 607 | proof | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 608 | assume ?P | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 609 | show ?Q | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 610 | proof | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 611 | fix a | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 612 | from \<open>?P\<close> have "min (count A a) (count B a) = 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 613 | by (simp add: multiset_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 614 | then have "count A a = 0 \<or> count B a = 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 615 | by (cases "count A a \<le> count B a") (simp_all add: min_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 616 | then show "a \<notin># A \<or> a \<notin># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 617 | by (simp add: not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 618 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 619 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 620 | assume ?Q | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 621 | show ?P | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 622 | proof (rule multiset_eqI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 623 | fix a | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 624 | from \<open>?Q\<close> have "count A a = 0 \<or> count B a = 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 625 | by (auto simp add: not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 626 |     then show "count (A #\<inter> B) a = count {#} a"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 627 | by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 628 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 629 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 630 | |
| 60606 | 631 | lemma empty_inter [simp]: "{#} #\<inter> M = {#}"
 | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 632 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 633 | |
| 60606 | 634 | lemma inter_empty [simp]: "M #\<inter> {#} = {#}"
 | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 635 | by (simp add: multiset_eq_iff) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 636 | |
| 60606 | 637 | lemma inter_add_left1: "\<not> x \<in># N \<Longrightarrow> (M + {#x#}) #\<inter> N = M #\<inter> N"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 638 | by (simp add: multiset_eq_iff not_in_iff) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 639 | |
| 60606 | 640 | lemma inter_add_left2: "x \<in># N \<Longrightarrow> (M + {#x#}) #\<inter> N = (M #\<inter> (N - {#x#})) + {#x#}"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 641 | by (auto simp add: multiset_eq_iff elim: mset_add) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 642 | |
| 60606 | 643 | lemma inter_add_right1: "\<not> x \<in># N \<Longrightarrow> N #\<inter> (M + {#x#}) = N #\<inter> M"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 644 | by (simp add: multiset_eq_iff not_in_iff) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 645 | |
| 60606 | 646 | lemma inter_add_right2: "x \<in># N \<Longrightarrow> N #\<inter> (M + {#x#}) = ((N - {#x#}) #\<inter> M) + {#x#}"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 647 | by (auto simp add: multiset_eq_iff elim: mset_add) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 648 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 649 | lemma disjunct_set_mset_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 650 |   assumes "M #\<inter> N = {#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 651 | shows "set_mset (M - N) = set_mset M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 652 | proof (rule set_eqI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 653 | fix a | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 654 | from assms have "a \<notin># M \<or> a \<notin># N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 655 | by (simp add: disjunct_not_in) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 656 | then show "a \<in># M - N \<longleftrightarrow> a \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 657 | by (auto dest: in_diffD) (simp add: in_diff_count not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 658 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 659 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 660 | lemma at_most_one_mset_mset_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 661 |   assumes "a \<notin># M - {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 662 |   shows "set_mset (M - {#a#}) = set_mset M - {a}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 663 | using assms by (auto simp add: not_in_iff in_diff_count set_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 664 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 665 | lemma more_than_one_mset_mset_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 666 |   assumes "a \<in># M - {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 667 |   shows "set_mset (M - {#a#}) = set_mset M"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 668 | proof (rule set_eqI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 669 | fix b | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 670 | have "Suc 0 < count M b \<Longrightarrow> count M b > 0" by arith | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 671 |   then show "b \<in># M - {#a#} \<longleftrightarrow> b \<in># M"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 672 | using assms by (auto simp add: in_diff_count) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 673 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 674 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 675 | lemma inter_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 676 | "a \<in># A #\<inter> B \<longleftrightarrow> a \<in># A \<and> a \<in># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 677 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 678 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 679 | lemma inter_union_distrib_left: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 680 | "A #\<inter> B + C = (A + C) #\<inter> (B + C)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 681 | by (simp add: multiset_eq_iff min_add_distrib_left) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 682 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 683 | lemma inter_union_distrib_right: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 684 | "C + A #\<inter> B = (C + A) #\<inter> (C + B)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 685 | using inter_union_distrib_left [of A B C] by (simp add: ac_simps) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 686 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 687 | lemma inter_subset_eq_union: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 688 | "A #\<inter> B \<subseteq># A + B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 689 | by (auto simp add: subseteq_mset_def) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 690 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 691 | |
| 60500 | 692 | subsubsection \<open>Bounded union\<close> | 
| 60678 | 693 | |
| 694 | definition sup_subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"(infixl "#\<union>" 70) | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 695 | where "sup_subset_mset A B = A + (B - A)" -- \<open>FIXME irregular fact name\<close> | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 696 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 697 | interpretation subset_mset: semilattice_sup sup_subset_mset "op \<subseteq>#" "op \<subset>#" | 
| 51623 | 698 | proof - | 
| 60678 | 699 | have [simp]: "m \<le> n \<Longrightarrow> q \<le> n \<Longrightarrow> m + (q - m) \<le> n" for m n q :: nat | 
| 700 | by arith | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 701 | show "class.semilattice_sup op #\<union> op \<subseteq># op \<subset>#" | 
| 60678 | 702 | by standard (auto simp add: sup_subset_mset_def subseteq_mset_def) | 
| 51623 | 703 | qed | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 704 | -- \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 705 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 706 | lemma sup_subset_mset_count [simp]: -- \<open>FIXME irregular fact name\<close> | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 707 | "count (A #\<union> B) x = max (count A x) (count B x)" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 708 | by (simp add: sup_subset_mset_def) | 
| 51623 | 709 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 710 | lemma set_mset_sup [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 711 | "set_mset (A #\<union> B) = set_mset A \<union> set_mset B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 712 | by (simp only: set_eq_iff count_greater_zero_iff [symmetric] sup_subset_mset_count) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 713 | (auto simp add: not_in_iff elim: mset_add) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 714 | |
| 60606 | 715 | lemma empty_sup [simp]: "{#} #\<union> M = M"
 | 
| 51623 | 716 | by (simp add: multiset_eq_iff) | 
| 717 | ||
| 60606 | 718 | lemma sup_empty [simp]: "M #\<union> {#} = M"
 | 
| 51623 | 719 | by (simp add: multiset_eq_iff) | 
| 720 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 721 | lemma sup_union_left1: "\<not> x \<in># N \<Longrightarrow> (M + {#x#}) #\<union> N = (M #\<union> N) + {#x#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 722 | by (simp add: multiset_eq_iff not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 723 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 724 | lemma sup_union_left2: "x \<in># N \<Longrightarrow> (M + {#x#}) #\<union> N = (M #\<union> (N - {#x#})) + {#x#}"
 | 
| 51623 | 725 | by (simp add: multiset_eq_iff) | 
| 726 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 727 | lemma sup_union_right1: "\<not> x \<in># N \<Longrightarrow> N #\<union> (M + {#x#}) = (N #\<union> M) + {#x#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 728 | by (simp add: multiset_eq_iff not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 729 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 730 | lemma sup_union_right2: "x \<in># N \<Longrightarrow> N #\<union> (M + {#x#}) = ((N - {#x#}) #\<union> M) + {#x#}"
 | 
| 51623 | 731 | by (simp add: multiset_eq_iff) | 
| 732 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 733 | lemma sup_union_distrib_left: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 734 | "A #\<union> B + C = (A + C) #\<union> (B + C)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 735 | by (simp add: multiset_eq_iff max_add_distrib_left) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 736 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 737 | lemma union_sup_distrib_right: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 738 | "C + A #\<union> B = (C + A) #\<union> (C + B)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 739 | using sup_union_distrib_left [of A B C] by (simp add: ac_simps) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 740 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 741 | lemma union_diff_inter_eq_sup: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 742 | "A + B - A #\<inter> B = A #\<union> B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 743 | by (auto simp add: multiset_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 744 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 745 | lemma union_diff_sup_eq_inter: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 746 | "A + B - A #\<union> B = A #\<inter> B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 747 | by (auto simp add: multiset_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 748 | |
| 51623 | 749 | |
| 60500 | 750 | subsubsection \<open>Subset is an order\<close> | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 751 | |
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 752 | interpretation subset_mset: order "op \<le>#" "op <#" by unfold_locales auto | 
| 51623 | 753 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 754 | |
| 60500 | 755 | subsubsection \<open>Filter (with comprehension syntax)\<close> | 
| 756 | ||
| 757 | text \<open>Multiset comprehension\<close> | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 758 | |
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 759 | lift_definition filter_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset"
 | 
| 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 760 | is "\<lambda>P M. \<lambda>x. if P x then M x else 0" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 761 | by (rule filter_preserves_multiset) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 762 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 763 | syntax (ASCII) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 764 |   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{# _ :# _./ _#})")
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 765 | syntax | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 766 |   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{# _ \<in># _./ _#})")
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 767 | translations | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 768 |   "{#x \<in># M. P#}" == "CONST filter_mset (\<lambda>x. P) M"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 769 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 770 | lemma count_filter_mset [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 771 | "count (filter_mset P M) a = (if P a then count M a else 0)" | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 772 | by (simp add: filter_mset.rep_eq) | 
| 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 773 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 774 | lemma set_mset_filter [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 775 |   "set_mset (filter_mset P M) = {a \<in> set_mset M. P a}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 776 | by (simp only: set_eq_iff count_greater_zero_iff [symmetric] count_filter_mset) simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 777 | |
| 60606 | 778 | lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}"
 | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 779 | by (rule multiset_eqI) simp | 
| 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 780 | |
| 60606 | 781 | lemma filter_single_mset [simp]: "filter_mset P {#x#} = (if P x then {#x#} else {#})"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 782 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 783 | |
| 60606 | 784 | lemma filter_union_mset [simp]: "filter_mset P (M + N) = filter_mset P M + filter_mset P N" | 
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 785 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 786 | |
| 60606 | 787 | lemma filter_diff_mset [simp]: "filter_mset P (M - N) = filter_mset P M - filter_mset P N" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 788 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 789 | |
| 60606 | 790 | lemma filter_inter_mset [simp]: "filter_mset P (M #\<inter> N) = filter_mset P M #\<inter> filter_mset P N" | 
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 791 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 792 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 793 | lemma multiset_filter_subset[simp]: "filter_mset f M \<subseteq># M" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 794 | by (simp add: mset_less_eqI) | 
| 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 795 | |
| 60606 | 796 | lemma multiset_filter_mono: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 797 | assumes "A \<subseteq># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 798 | shows "filter_mset f A \<subseteq># filter_mset f B" | 
| 58035 | 799 | proof - | 
| 800 | from assms[unfolded mset_le_exists_conv] | |
| 801 | obtain C where B: "B = A + C" by auto | |
| 802 | show ?thesis unfolding B by auto | |
| 803 | qed | |
| 804 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 805 | lemma filter_mset_eq_conv: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 806 | "filter_mset P M = N \<longleftrightarrow> N \<subseteq># M \<and> (\<forall>b\<in>#N. P b) \<and> (\<forall>a\<in>#M - N. \<not> P a)" (is "?P \<longleftrightarrow> ?Q") | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 807 | proof | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 808 | assume ?P then show ?Q by auto (simp add: multiset_eq_iff in_diff_count) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 809 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 810 | assume ?Q | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 811 | then obtain Q where M: "M = N + Q" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 812 | by (auto simp add: mset_le_exists_conv) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 813 | then have MN: "M - N = Q" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 814 | show ?P | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 815 | proof (rule multiset_eqI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 816 | fix a | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 817 | from \<open>?Q\<close> MN have *: "\<not> P a \<Longrightarrow> a \<notin># N" "P a \<Longrightarrow> a \<notin># Q" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 818 | by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 819 | show "count (filter_mset P M) a = count N a" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 820 | proof (cases "a \<in># M") | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 821 | case True | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 822 | with * show ?thesis | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 823 | by (simp add: not_in_iff M) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 824 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 825 | case False then have "count M a = 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 826 | by (simp add: not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 827 | with M show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 828 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 829 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 830 | qed | 
| 59813 | 831 | |
| 832 | ||
| 60500 | 833 | subsubsection \<open>Size\<close> | 
| 10249 | 834 | |
| 56656 | 835 | definition wcount where "wcount f M = (\<lambda>x. count M x * Suc (f x))" | 
| 836 | ||
| 837 | lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" | |
| 838 | by (auto simp: wcount_def add_mult_distrib) | |
| 839 | ||
| 840 | definition size_multiset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a multiset \<Rightarrow> nat" where
 | |
| 60495 | 841 | "size_multiset f M = setsum (wcount f M) (set_mset M)" | 
| 56656 | 842 | |
| 843 | lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] | |
| 844 | ||
| 60606 | 845 | instantiation multiset :: (type) size | 
| 846 | begin | |
| 847 | ||
| 56656 | 848 | definition size_multiset where | 
| 849 | 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 | 850 | instance .. | 
| 60606 | 851 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 852 | 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 | 853 | |
| 56656 | 854 | lemmas size_multiset_overloaded_eq = | 
| 855 | size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] | |
| 856 | ||
| 857 | lemma size_multiset_empty [simp]: "size_multiset f {#} = 0"
 | |
| 858 | by (simp add: size_multiset_def) | |
| 859 | ||
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 860 | lemma size_empty [simp]: "size {#} = 0"
 | 
| 56656 | 861 | by (simp add: size_multiset_overloaded_def) | 
| 862 | ||
| 863 | lemma size_multiset_single [simp]: "size_multiset f {#b#} = Suc (f b)"
 | |
| 864 | by (simp add: size_multiset_eq) | |
| 10249 | 865 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 866 | lemma size_single [simp]: "size {#b#} = 1"
 | 
| 56656 | 867 | by (simp add: size_multiset_overloaded_def) | 
| 868 | ||
| 869 | lemma setsum_wcount_Int: | |
| 60495 | 870 | "finite A \<Longrightarrow> setsum (wcount f N) (A \<inter> set_mset N) = setsum (wcount f N) A" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 871 | by (induct rule: finite_induct) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 872 | (simp_all add: Int_insert_left wcount_def count_eq_zero_iff) | 
| 56656 | 873 | |
| 874 | lemma size_multiset_union [simp]: | |
| 875 | "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" | |
| 57418 | 876 | apply (simp add: size_multiset_def setsum_Un_nat setsum.distrib setsum_wcount_Int wcount_union) | 
| 56656 | 877 | apply (subst Int_commute) | 
| 878 | apply (simp add: setsum_wcount_Int) | |
| 26178 | 879 | done | 
| 10249 | 880 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 881 | lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" | 
| 56656 | 882 | by (auto simp add: size_multiset_overloaded_def) | 
| 883 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 884 | lemma size_multiset_eq_0_iff_empty [iff]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 885 |   "size_multiset f M = 0 \<longleftrightarrow> M = {#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 886 | by (auto simp add: size_multiset_eq count_eq_zero_iff) | 
| 10249 | 887 | |
| 17161 | 888 | lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
 | 
| 56656 | 889 | by (auto simp add: size_multiset_overloaded_def) | 
| 26016 | 890 | |
| 891 | lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
 | |
| 26178 | 892 | by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) | 
| 10249 | 893 | |
| 60607 | 894 | lemma size_eq_Suc_imp_elem: "size M = Suc n \<Longrightarrow> \<exists>a. a \<in># M" | 
| 56656 | 895 | apply (unfold size_multiset_overloaded_eq) | 
| 26178 | 896 | apply (drule setsum_SucD) | 
| 897 | apply auto | |
| 898 | done | |
| 10249 | 899 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 900 | 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 | 901 | 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 | 902 |   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 | 903 | 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 | 904 | 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 | 905 | 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 | 906 |   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 | 907 | then show ?thesis by blast | 
| 23611 | 908 | qed | 
| 15869 | 909 | |
| 60606 | 910 | lemma size_mset_mono: | 
| 911 | fixes A B :: "'a multiset" | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 912 | assumes "A \<subseteq># B" | 
| 60606 | 913 | shows "size A \<le> size B" | 
| 59949 | 914 | proof - | 
| 915 | from assms[unfolded mset_le_exists_conv] | |
| 916 | obtain C where B: "B = A + C" by auto | |
| 60606 | 917 | show ?thesis unfolding B by (induct C) auto | 
| 59949 | 918 | qed | 
| 919 | ||
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 920 | lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \<le> size M" | 
| 59949 | 921 | by (rule size_mset_mono[OF multiset_filter_subset]) | 
| 922 | ||
| 923 | lemma size_Diff_submset: | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 924 | "M \<subseteq># M' \<Longrightarrow> size (M' - M) = size M' - size(M::'a multiset)" | 
| 59949 | 925 | by (metis add_diff_cancel_left' size_union mset_le_exists_conv) | 
| 26016 | 926 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 927 | |
| 60500 | 928 | subsection \<open>Induction and case splits\<close> | 
| 10249 | 929 | |
| 18258 | 930 | theorem multiset_induct [case_names empty add, induct type: multiset]: | 
| 48009 | 931 |   assumes empty: "P {#}"
 | 
| 932 |   assumes add: "\<And>M x. P M \<Longrightarrow> P (M + {#x#})"
 | |
| 933 | shows "P M" | |
| 934 | proof (induct n \<equiv> "size M" arbitrary: M) | |
| 935 | case 0 thus "P M" by (simp add: empty) | |
| 936 | next | |
| 937 | case (Suc k) | |
| 938 |   obtain N x where "M = N + {#x#}"
 | |
| 60500 | 939 | using \<open>Suc k = size M\<close> [symmetric] | 
| 48009 | 940 | using size_eq_Suc_imp_eq_union by fast | 
| 941 | with Suc add show "P M" by simp | |
| 10249 | 942 | qed | 
| 943 | ||
| 25610 | 944 | lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = A + {#a#}"
 | 
| 26178 | 945 | by (induct M) auto | 
| 25610 | 946 | |
| 55913 | 947 | lemma multiset_cases [cases type]: | 
| 948 |   obtains (empty) "M = {#}"
 | |
| 949 |     | (add) N x where "M = N + {#x#}"
 | |
| 950 | using assms by (induct M) simp_all | |
| 25610 | 951 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | 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 | 953 | 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 | 954 | |
| 60607 | 955 | lemma multiset_partition: "M = {# x\<in>#M. P x #} + {# x\<in>#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 | 956 | apply (subst multiset_eq_iff) | 
| 26178 | 957 | apply auto | 
| 958 | done | |
| 10249 | 959 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 960 | lemma mset_less_size: "(A::'a multiset) \<subset># B \<Longrightarrow> size A < size B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 961 | 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 | 962 | 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 | 963 |   then have "M \<noteq> {#}" by (simp add: mset_less_empty_nonempty)
 | 
| 58425 | 964 |   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 | 965 | 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 | 966 | 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 | 967 | 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 | 968 | case (add S x T) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 969 | have IH: "\<And>B. S \<subset># B \<Longrightarrow> size S < size B" by fact | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 970 |   have SxsubT: "S + {#x#} \<subset># T" by fact
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 971 | then have "x \<in># T" and "S \<subset># T" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 972 | by (auto dest: mset_less_insertD) | 
| 58425 | 973 |   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 | 974 | by (blast dest: multi_member_split) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 975 | then have "S \<subset># 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 | 976 | 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 | 977 | 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 | 978 | 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 | 979 | 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 | 980 | |
| 59949 | 981 | lemma size_1_singleton_mset: "size M = 1 \<Longrightarrow> \<exists>a. M = {#a#}"
 | 
| 982 | by (cases M) auto | |
| 983 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 984 | |
| 60500 | 985 | subsubsection \<open>Strong induction and subset induction for multisets\<close> | 
| 986 | ||
| 987 | text \<open>Well-foundedness of strict subset relation\<close> | |
| 58098 | 988 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 989 | lemma wf_less_mset_rel: "wf {(M, N :: 'a multiset). M \<subset># N}"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 990 | 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 | 991 | 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 | 992 | 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 | 993 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma full_multiset_induct [case_names less]: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 995 | assumes ih: "\<And>B. \<forall>(A::'a multiset). A \<subset># B \<longrightarrow> P A \<Longrightarrow> P B" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 996 | shows "P B" | 
| 58098 | 997 | apply (rule wf_less_mset_rel [THEN wf_induct]) | 
| 998 | 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 | 999 | 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 | 1000 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned 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 | lemma multi_subset_induct [consumes 2, case_names empty add]: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1002 | assumes "F \<subseteq># A" | 
| 60606 | 1003 |     and empty: "P {#}"
 | 
| 1004 |     and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (F + {#a#})"
 | |
| 1005 | shows "P F" | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1006 | proof - | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1007 | from \<open>F \<subseteq># A\<close> | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1008 | 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 | 1009 | 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 | 1010 |     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 | 1011 | 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 | 1012 | fix x F | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1013 |     assume P: "F \<subseteq># A \<Longrightarrow> P F" and i: "F + {#x#} \<subseteq># 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 | 1014 |     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 | 1015 | 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 | 1016 | from i show "x \<in># A" by (auto dest: mset_le_insertD) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1017 | from i have "F \<subseteq># 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 | 1018 | 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 | 1019 | 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 | 1020 | 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 | 1021 | qed | 
| 26145 | 1022 | |
| 17161 | 1023 | |
| 60500 | 1024 | subsection \<open>The fold combinator\<close> | 
| 48023 | 1025 | |
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1026 | definition fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b"
 | 
| 48023 | 1027 | where | 
| 60495 | 1028 | "fold_mset f s M = Finite_Set.fold (\<lambda>x. f x ^^ count M x) s (set_mset M)" | 
| 48023 | 1029 | |
| 60606 | 1030 | lemma fold_mset_empty [simp]: "fold_mset f s {#} = s"
 | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1031 | by (simp add: fold_mset_def) | 
| 48023 | 1032 | |
| 1033 | context comp_fun_commute | |
| 1034 | begin | |
| 1035 | ||
| 60606 | 1036 | lemma fold_mset_insert: "fold_mset f s (M + {#x#}) = f x (fold_mset f s M)"
 | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1037 | proof - | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1038 | interpret mset: comp_fun_commute "\<lambda>y. f y ^^ count M y" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1039 | by (fact comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1040 |   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 | 1041 | by (fact comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1042 | show ?thesis | 
| 60495 | 1043 | proof (cases "x \<in> set_mset M") | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1044 | case False | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1045 |     then have *: "count (M + {#x#}) x = 1"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1046 | by (simp add: not_in_iff) | 
| 60495 | 1047 |     from False have "Finite_Set.fold (\<lambda>y. f y ^^ count (M + {#x#}) y) s (set_mset M) =
 | 
| 1048 | Finite_Set.fold (\<lambda>y. f y ^^ count M y) s (set_mset M)" | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1049 | by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1050 | with False * show ?thesis | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1051 | by (simp add: fold_mset_def del: count_union) | 
| 48023 | 1052 | next | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1053 | case True | 
| 60495 | 1054 |     def N \<equiv> "set_mset M - {x}"
 | 
| 1055 | from N_def True have *: "set_mset M = insert x N" "x \<notin> N" "finite N" by auto | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1056 |     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 | 1057 | Finite_Set.fold (\<lambda>y. f y ^^ count M y) s N" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1058 | by (auto intro!: Finite_Set.fold_cong comp_fun_commute_funpow) | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1059 | with * show ?thesis by (simp add: fold_mset_def del: count_union) simp | 
| 48023 | 1060 | qed | 
| 1061 | qed | |
| 1062 | ||
| 60606 | 1063 | corollary fold_mset_single [simp]: "fold_mset f s {#x#} = f x s"
 | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1064 | proof - | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1065 |   have "fold_mset f s ({#} + {#x#}) = f x s" by (simp only: fold_mset_insert) simp
 | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1066 | then show ?thesis by simp | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1067 | qed | 
| 48023 | 1068 | |
| 60606 | 1069 | lemma fold_mset_fun_left_comm: "f x (fold_mset f s M) = fold_mset f (f x s) M" | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1070 | by (induct M) (simp_all add: fold_mset_insert fun_left_comm) | 
| 48023 | 1071 | |
| 60606 | 1072 | lemma fold_mset_union [simp]: "fold_mset f s (M + N) = fold_mset f (fold_mset f s M) N" | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1073 | proof (induct M) | 
| 48023 | 1074 | case empty then show ?case by simp | 
| 1075 | next | |
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1076 | case (add M x) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1077 |   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 | 1078 | by (simp add: ac_simps) | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1079 | with add show ?case by (simp add: fold_mset_insert fold_mset_fun_left_comm) | 
| 48023 | 1080 | qed | 
| 1081 | ||
| 1082 | lemma fold_mset_fusion: | |
| 1083 | assumes "comp_fun_commute g" | |
| 60606 | 1084 | and *: "\<And>x y. h (g x y) = f x (h y)" | 
| 1085 | shows "h (fold_mset g w A) = fold_mset f (h w) A" | |
| 48023 | 1086 | proof - | 
| 1087 | interpret comp_fun_commute g by (fact assms) | |
| 60606 | 1088 | from * show ?thesis by (induct A) auto | 
| 48023 | 1089 | qed | 
| 1090 | ||
| 1091 | end | |
| 1092 | ||
| 60500 | 1093 | text \<open> | 
| 48023 | 1094 | A note on code generation: When defining some function containing a | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1095 |   subterm @{term "fold_mset F"}, code generation is not automatic. When
 | 
| 61585 | 1096 | interpreting locale \<open>left_commutative\<close> with \<open>F\<close>, the | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1097 |   would be code thms for @{const fold_mset} become thms like
 | 
| 61585 | 1098 |   @{term "fold_mset F z {#} = z"} where \<open>F\<close> is not a pattern but
 | 
| 48023 | 1099 | contains defined symbols, i.e.\ is not a code thm. Hence a separate | 
| 61585 | 1100 | constant with its own code thms needs to be introduced for \<open>F\<close>. See the image operator below. | 
| 60500 | 1101 | \<close> | 
| 1102 | ||
| 1103 | ||
| 1104 | subsection \<open>Image\<close> | |
| 48023 | 1105 | |
| 1106 | definition image_mset :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
 | |
| 60607 | 1107 |   "image_mset f = fold_mset (plus \<circ> single \<circ> f) {#}"
 | 
| 1108 | ||
| 1109 | lemma comp_fun_commute_mset_image: "comp_fun_commute (plus \<circ> single \<circ> f)" | |
| 49823 | 1110 | proof | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1111 | qed (simp add: ac_simps fun_eq_iff) | 
| 48023 | 1112 | |
| 1113 | lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
 | |
| 49823 | 1114 | by (simp add: image_mset_def) | 
| 48023 | 1115 | |
| 1116 | lemma image_mset_single [simp]: "image_mset f {#x#} = {#f x#}"
 | |
| 49823 | 1117 | proof - | 
| 60607 | 1118 | interpret comp_fun_commute "plus \<circ> single \<circ> f" | 
| 49823 | 1119 | by (fact comp_fun_commute_mset_image) | 
| 1120 | show ?thesis by (simp add: image_mset_def) | |
| 1121 | qed | |
| 48023 | 1122 | |
| 60606 | 1123 | lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" | 
| 49823 | 1124 | proof - | 
| 60607 | 1125 | interpret comp_fun_commute "plus \<circ> single \<circ> f" | 
| 49823 | 1126 | 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 | 1127 | show ?thesis by (induct N) (simp_all add: image_mset_def ac_simps) | 
| 49823 | 1128 | qed | 
| 1129 | ||
| 60606 | 1130 | corollary image_mset_insert: "image_mset f (M + {#a#}) = image_mset f M + {#f a#}"
 | 
| 49823 | 1131 | by simp | 
| 48023 | 1132 | |
| 60606 | 1133 | lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)" | 
| 49823 | 1134 | by (induct M) simp_all | 
| 48040 | 1135 | |
| 60606 | 1136 | lemma size_image_mset [simp]: "size (image_mset f M) = size M" | 
| 49823 | 1137 | by (induct M) simp_all | 
| 48023 | 1138 | |
| 60606 | 1139 | lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \<longleftrightarrow> M = {#}"
 | 
| 49823 | 1140 | by (cases M) auto | 
| 48023 | 1141 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1142 | syntax (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1143 |   "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ :# _#})")
 | 
| 48023 | 1144 | syntax | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1145 |   "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ \<in># _#})")
 | 
| 59813 | 1146 | translations | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1147 |   "{#e. x \<in># M#}" \<rightleftharpoons> "CONST image_mset (\<lambda>x. e) M"
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1148 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1149 | syntax (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1150 |   "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ :# _./ _#})")
 | 
| 48023 | 1151 | syntax | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1152 |   "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ \<in># _./ _#})")
 | 
| 59813 | 1153 | translations | 
| 60606 | 1154 |   "{#e | x\<in>#M. P#}" \<rightharpoonup> "{#e. x \<in># {# x\<in>#M. P#}#}"
 | 
| 59813 | 1155 | |
| 60500 | 1156 | text \<open> | 
| 60607 | 1157 |   This allows to write not just filters like @{term "{#x\<in>#M. x<c#}"}
 | 
| 1158 |   but also images like @{term "{#x+x. x\<in>#M #}"} and @{term [source]
 | |
| 1159 |   "{#x+x|x\<in>#M. x<c#}"}, where the latter is currently displayed as
 | |
| 1160 |   @{term "{#x+x|x\<in>#M. x<c#}"}.
 | |
| 60500 | 1161 | \<close> | 
| 48023 | 1162 | |
| 60495 | 1163 | lemma in_image_mset: "y \<in># {#f x. x \<in># M#} \<longleftrightarrow> y \<in> f ` set_mset M"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1164 | by (metis set_image_mset) | 
| 59813 | 1165 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55417diff
changeset | 1166 | functor image_mset: image_mset | 
| 48023 | 1167 | proof - | 
| 1168 | fix f g show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)" | |
| 1169 | proof | |
| 1170 | fix A | |
| 1171 | show "(image_mset f \<circ> image_mset g) A = image_mset (f \<circ> g) A" | |
| 1172 | by (induct A) simp_all | |
| 1173 | qed | |
| 1174 | show "image_mset id = id" | |
| 1175 | proof | |
| 1176 | fix A | |
| 1177 | show "image_mset id A = id A" | |
| 1178 | by (induct A) simp_all | |
| 1179 | qed | |
| 1180 | qed | |
| 1181 | ||
| 59813 | 1182 | declare | 
| 1183 | image_mset.id [simp] | |
| 1184 | image_mset.identity [simp] | |
| 1185 | ||
| 1186 | lemma image_mset_id[simp]: "image_mset id x = x" | |
| 1187 | unfolding id_def by auto | |
| 1188 | ||
| 1189 | lemma image_mset_cong: "(\<And>x. x \<in># M \<Longrightarrow> f x = g x) \<Longrightarrow> {#f x. x \<in># M#} = {#g x. x \<in># M#}"
 | |
| 1190 | by (induct M) auto | |
| 1191 | ||
| 1192 | lemma image_mset_cong_pair: | |
| 1193 |   "(\<forall>x y. (x, y) \<in># M \<longrightarrow> f x y = g x y) \<Longrightarrow> {#f x y. (x, y) \<in># M#} = {#g x y. (x, y) \<in># M#}"
 | |
| 1194 | by (metis image_mset_cong split_cong) | |
| 49717 | 1195 | |
| 48023 | 1196 | |
| 60500 | 1197 | subsection \<open>Further conversions\<close> | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1198 | |
| 60515 | 1199 | primrec mset :: "'a list \<Rightarrow> 'a multiset" where | 
| 1200 |   "mset [] = {#}" |
 | |
| 1201 |   "mset (a # x) = mset x + {# 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 | 1202 | |
| 37107 | 1203 | lemma in_multiset_in_set: | 
| 60515 | 1204 | "x \<in># mset xs \<longleftrightarrow> x \<in> set xs" | 
| 37107 | 1205 | by (induct xs) simp_all | 
| 1206 | ||
| 60515 | 1207 | lemma count_mset: | 
| 1208 | "count (mset xs) x = length (filter (\<lambda>y. x = y) xs)" | |
| 37107 | 1209 | by (induct xs) simp_all | 
| 1210 | ||
| 60515 | 1211 | lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])"
 | 
| 59813 | 1212 | by (induct x) auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1213 | |
| 60515 | 1214 | lemma mset_zero_iff_right[simp]: "({#} = mset x) = (x = [])"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1215 | 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 | 1216 | |
| 60515 | 1217 | lemma set_mset_mset[simp]: "set_mset (mset 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 | 1218 | 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 | 1219 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1220 | lemma set_mset_comp_mset [simp]: "set_mset \<circ> mset = set" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1221 | by (simp add: fun_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1222 | |
| 60515 | 1223 | lemma size_mset [simp]: "size (mset xs) = length xs" | 
| 48012 | 1224 | by (induct xs) simp_all | 
| 1225 | ||
| 60606 | 1226 | lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1227 | 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 | 1228 | |
| 60607 | 1229 | lemma mset_filter: "mset (filter P xs) = {#x \<in># mset xs. P x #}"
 | 
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1230 | by (induct xs) simp_all | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1231 | |
| 60515 | 1232 | lemma mset_rev [simp]: | 
| 1233 | "mset (rev xs) = mset xs" | |
| 40950 | 1234 | by (induct xs) simp_all | 
| 1235 | ||
| 60515 | 1236 | lemma surj_mset: "surj mset" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1237 | 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 | 1238 | 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 | 1239 | 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 | 1240 | 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 | 1241 | 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 | 1242 | 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 | 1243 | 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 | 1244 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1245 | lemma distinct_count_atmost_1: | 
| 60606 | 1246 | "distinct x = (\<forall>a. count (mset x) a = (if a \<in> set x then 1 else 0))" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1247 | proof (induct x) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1248 | case Nil then show ?case by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1249 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1250 | case (Cons x xs) show ?case (is "?lhs \<longleftrightarrow> ?rhs") | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1251 | proof | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1252 | assume ?lhs then show ?rhs using Cons by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1253 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1254 | assume ?rhs then have "x \<notin> set xs" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1255 | by (simp split: if_splits) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1256 | moreover from \<open>?rhs\<close> have "(\<forall>a. count (mset xs) a = | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1257 | (if a \<in> set xs then 1 else 0))" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1258 | by (auto split: if_splits simp add: count_eq_zero_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1259 | ultimately show ?lhs using Cons by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1260 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1261 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1262 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1263 | lemma mset_eq_setD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1264 | assumes "mset xs = mset ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1265 | shows "set xs = set ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1266 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1267 | from assms have "set_mset (mset xs) = set_mset (mset ys)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1268 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1269 | then show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1270 | qed | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1271 | |
| 60515 | 1272 | lemma set_eq_iff_mset_eq_distinct: | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1273 | "distinct x \<Longrightarrow> distinct y \<Longrightarrow> | 
| 60515 | 1274 | (set x = set y) = (mset x = mset y)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1275 | 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 | 1276 | |
| 60515 | 1277 | lemma set_eq_iff_mset_remdups_eq: | 
| 1278 | "(set x = set y) = (mset (remdups x) = mset (remdups y))" | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1279 | apply (rule iffI) | 
| 60515 | 1280 | apply (simp add: set_eq_iff_mset_eq_distinct[THEN iffD1]) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1281 | apply (drule distinct_remdups [THEN distinct_remdups | 
| 60515 | 1282 | [THEN set_eq_iff_mset_eq_distinct [THEN iffD2]]]) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1283 | 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 | 1284 | 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 | 1285 | |
| 60606 | 1286 | lemma mset_compl_union [simp]: "mset [x\<leftarrow>xs. P x] + mset [x\<leftarrow>xs. \<not>P x] = mset xs" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 1287 | 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 | 1288 | |
| 60607 | 1289 | lemma nth_mem_mset: "i < length ls \<Longrightarrow> (ls ! i) \<in># mset ls" | 
| 60678 | 1290 | proof (induct ls arbitrary: i) | 
| 1291 | case Nil | |
| 1292 | then show ?case by simp | |
| 1293 | next | |
| 1294 | case Cons | |
| 1295 | then show ?case by (cases i) auto | |
| 1296 | qed | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1297 | |
| 60606 | 1298 | lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}"
 | 
| 60678 | 1299 | 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 | 1300 | |
| 60515 | 1301 | lemma mset_eq_length: | 
| 1302 | assumes "mset xs = mset ys" | |
| 37107 | 1303 | shows "length xs = length ys" | 
| 60515 | 1304 | using assms by (metis size_mset) | 
| 1305 | ||
| 1306 | lemma mset_eq_length_filter: | |
| 1307 | assumes "mset xs = mset ys" | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1308 | shows "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) ys)" | 
| 60515 | 1309 | using assms by (metis count_mset) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1310 | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1311 | lemma fold_multiset_equiv: | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1312 | 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" | 
| 60515 | 1313 | and equiv: "mset xs = mset ys" | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1314 | shows "List.fold f xs = List.fold f ys" | 
| 60606 | 1315 | using f equiv [symmetric] | 
| 46921 | 1316 | proof (induct xs arbitrary: ys) | 
| 60678 | 1317 | case Nil | 
| 1318 | then show ?case by simp | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1319 | next | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1320 | case (Cons x xs) | 
| 60678 | 1321 | then have *: "set ys = set (x # xs)" | 
| 1322 | by (blast dest: mset_eq_setD) | |
| 58425 | 1323 | 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 | 1324 | by (rule Cons.prems(1)) (simp_all add: *) | 
| 60678 | 1325 | moreover from * have "x \<in> set ys" | 
| 1326 | by simp | |
| 1327 | ultimately have "List.fold f ys = List.fold f (remove1 x ys) \<circ> f x" | |
| 1328 | by (fact fold_remove1_split) | |
| 1329 | moreover from Cons.prems have "List.fold f xs = List.fold f (remove1 x ys)" | |
| 1330 | by (auto intro: Cons.hyps) | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1331 | ultimately show ?case by simp | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1332 | qed | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 1333 | |
| 60606 | 1334 | lemma mset_insort [simp]: "mset (insort x xs) = mset xs + {#x#}"
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1335 | by (induct xs) (simp_all add: ac_simps) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1336 | |
| 60606 | 1337 | lemma mset_map: "mset (map f xs) = image_mset f (mset xs)" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1338 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1339 | |
| 61890 
f6ded81f5690
abandoned attempt to unify sublocale and interpretation into global theories
 haftmann parents: 
61832diff
changeset | 1340 | global_interpretation mset_set: folding "\<lambda>x M. {#x#} + M" "{#}"
 | 
| 61832 | 1341 |   defines mset_set = "folding.F (\<lambda>x M. {#x#} + M) {#}"
 | 
| 1342 | by standard (simp add: fun_eq_iff ac_simps) | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1343 | |
| 60513 | 1344 | lemma count_mset_set [simp]: | 
| 1345 | "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> count (mset_set A) x = 1" (is "PROP ?P") | |
| 1346 | "\<not> finite A \<Longrightarrow> count (mset_set A) x = 0" (is "PROP ?Q") | |
| 1347 | "x \<notin> A \<Longrightarrow> count (mset_set A) x = 0" (is "PROP ?R") | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1348 | proof - | 
| 60606 | 1349 | have *: "count (mset_set A) x = 0" if "x \<notin> A" for A | 
| 1350 | proof (cases "finite A") | |
| 1351 | case False then show ?thesis by simp | |
| 1352 | next | |
| 1353 | case True from True \<open>x \<notin> A\<close> show ?thesis by (induct A) auto | |
| 1354 | qed | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 1355 | 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 | 1356 | by (auto elim!: Set.set_insert) | 
| 61585 | 1357 | qed \<comment> \<open>TODO: maybe define @{const mset_set} also in terms of @{const Abs_multiset}\<close>
 | 
| 60513 | 1358 | |
| 1359 | lemma elem_mset_set[simp, intro]: "finite A \<Longrightarrow> x \<in># mset_set A \<longleftrightarrow> x \<in> A" | |
| 59813 | 1360 | by (induct A rule: finite_induct) simp_all | 
| 1361 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1362 | context linorder | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1363 | begin | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1364 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1365 | definition sorted_list_of_multiset :: "'a multiset \<Rightarrow> 'a list" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1366 | where | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1367 | "sorted_list_of_multiset M = fold_mset insort [] M" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1368 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1369 | lemma sorted_list_of_multiset_empty [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1370 |   "sorted_list_of_multiset {#} = []"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1371 | by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1372 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1373 | lemma sorted_list_of_multiset_singleton [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1374 |   "sorted_list_of_multiset {#x#} = [x]"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1375 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1376 | interpret comp_fun_commute insort by (fact comp_fun_commute_insort) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1377 | show ?thesis by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1378 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1379 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1380 | lemma sorted_list_of_multiset_insert [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1381 |   "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 | 1382 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1383 | interpret comp_fun_commute insort by (fact comp_fun_commute_insort) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1384 | show ?thesis by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1385 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1386 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1387 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1388 | |
| 60515 | 1389 | lemma mset_sorted_list_of_multiset [simp]: | 
| 1390 | "mset (sorted_list_of_multiset M) = M" | |
| 60513 | 1391 | by (induct M) simp_all | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1392 | |
| 60515 | 1393 | lemma sorted_list_of_multiset_mset [simp]: | 
| 1394 | "sorted_list_of_multiset (mset xs) = sort xs" | |
| 60513 | 1395 | by (induct xs) simp_all | 
| 1396 | ||
| 1397 | lemma finite_set_mset_mset_set[simp]: | |
| 1398 | "finite A \<Longrightarrow> set_mset (mset_set A) = A" | |
| 1399 | by (induct A rule: finite_induct) simp_all | |
| 1400 | ||
| 1401 | lemma infinite_set_mset_mset_set: | |
| 1402 |   "\<not> finite A \<Longrightarrow> set_mset (mset_set A) = {}"
 | |
| 1403 | by simp | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1404 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1405 | lemma set_sorted_list_of_multiset [simp]: | 
| 60495 | 1406 | "set (sorted_list_of_multiset M) = set_mset M" | 
| 60513 | 1407 | by (induct M) (simp_all add: set_insort) | 
| 1408 | ||
| 1409 | lemma sorted_list_of_mset_set [simp]: | |
| 1410 | "sorted_list_of_multiset (mset_set A) = sorted_list_of_set A" | |
| 1411 | by (cases "finite A") (induct A rule: finite_induct, simp_all add: ac_simps) | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1412 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1413 | |
| 60804 | 1414 | subsection \<open>Replicate operation\<close> | 
| 1415 | ||
| 1416 | definition replicate_mset :: "nat \<Rightarrow> 'a \<Rightarrow> 'a multiset" where | |
| 1417 |   "replicate_mset n x = ((op + {#x#}) ^^ n) {#}"
 | |
| 1418 | ||
| 1419 | lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}"
 | |
| 1420 | unfolding replicate_mset_def by simp | |
| 1421 | ||
| 1422 | lemma replicate_mset_Suc[simp]: "replicate_mset (Suc n) x = replicate_mset n x + {#x#}"
 | |
| 1423 | unfolding replicate_mset_def by (induct n) (auto intro: add.commute) | |
| 1424 | ||
| 1425 | lemma in_replicate_mset[simp]: "x \<in># replicate_mset n y \<longleftrightarrow> n > 0 \<and> x = y" | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1426 | unfolding replicate_mset_def by (induct n) auto | 
| 60804 | 1427 | |
| 1428 | lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)" | |
| 1429 | unfolding replicate_mset_def by (induct n) simp_all | |
| 1430 | ||
| 1431 | lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})"
 | |
| 1432 | by (auto split: if_splits) | |
| 1433 | ||
| 1434 | lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" | |
| 1435 | by (induct n, simp_all) | |
| 1436 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1437 | lemma count_le_replicate_mset_le: "n \<le> count M x \<longleftrightarrow> replicate_mset n x \<subseteq># M" | 
| 60804 | 1438 | by (auto simp add: assms mset_less_eqI) (metis count_replicate_mset subseteq_mset_def) | 
| 1439 | ||
| 1440 | lemma filter_eq_replicate_mset: "{#y \<in># D. y = x#} = replicate_mset (count D x) x"
 | |
| 1441 | by (induct D) simp_all | |
| 1442 | ||
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1443 | lemma replicate_count_mset_eq_filter_eq: | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1444 | "replicate (count (mset xs) k) k = filter (HOL.eq k) xs" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1445 | by (induct xs) auto | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1446 | |
| 62366 | 1447 | lemma replicate_mset_eq_empty_iff [simp]: | 
| 1448 |   "replicate_mset n a = {#} \<longleftrightarrow> n = 0"
 | |
| 1449 | by (induct n) simp_all | |
| 1450 | ||
| 1451 | lemma replicate_mset_eq_iff: | |
| 1452 | "replicate_mset m a = replicate_mset n b \<longleftrightarrow> | |
| 1453 | m = 0 \<and> n = 0 \<or> m = n \<and> a = b" | |
| 1454 | by (auto simp add: multiset_eq_iff) | |
| 1455 | ||
| 60804 | 1456 | |
| 60500 | 1457 | subsection \<open>Big operators\<close> | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1458 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1459 | no_notation times (infixl "*" 70) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1460 | no_notation Groups.one ("1")
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1461 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1462 | locale comm_monoid_mset = comm_monoid | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1463 | begin | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1464 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1465 | definition F :: "'a multiset \<Rightarrow> 'a" | 
| 60606 | 1466 | where eq_fold: "F M = fold_mset f 1 M" | 
| 1467 | ||
| 1468 | lemma empty [simp]: "F {#} = 1"
 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1469 | by (simp add: eq_fold) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1470 | |
| 60678 | 1471 | lemma singleton [simp]: "F {#x#} = x"
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1472 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1473 | interpret comp_fun_commute | 
| 60678 | 1474 | by standard (simp add: fun_eq_iff left_commute) | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1475 | show ?thesis by (simp add: eq_fold) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1476 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1477 | |
| 60606 | 1478 | lemma union [simp]: "F (M + N) = F M * F N" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1479 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1480 | interpret comp_fun_commute f | 
| 60678 | 1481 | by standard (simp add: fun_eq_iff left_commute) | 
| 1482 | show ?thesis | |
| 1483 | by (induct N) (simp_all add: left_commute eq_fold) | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1484 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1485 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1486 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1487 | |
| 61076 | 1488 | lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute (op + :: 'a multiset \<Rightarrow> _ \<Rightarrow> _)" | 
| 60678 | 1489 | by standard (simp add: add_ac comp_def) | 
| 59813 | 1490 | |
| 1491 | declare comp_fun_commute.fold_mset_insert[OF comp_fun_commute_plus_mset, simp] | |
| 1492 | ||
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1493 | lemma in_mset_fold_plus_iff[iff]: "x \<in># fold_mset (op +) M NN \<longleftrightarrow> x \<in># M \<or> (\<exists>N. N \<in># NN \<and> x \<in># N)" | 
| 59813 | 1494 | by (induct NN) auto | 
| 1495 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1496 | notation times (infixl "*" 70) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1497 | notation Groups.one ("1")
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1498 | |
| 54868 | 1499 | context comm_monoid_add | 
| 1500 | begin | |
| 1501 | ||
| 61605 | 1502 | sublocale msetsum: comm_monoid_mset plus 0 | 
| 61832 | 1503 | defines msetsum = msetsum.F .. | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1504 | |
| 60804 | 1505 | lemma (in semiring_1) msetsum_replicate_mset [simp]: | 
| 1506 | "msetsum (replicate_mset n a) = of_nat n * a" | |
| 1507 | by (induct n) (simp_all add: algebra_simps) | |
| 1508 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1509 | lemma setsum_unfold_msetsum: | 
| 60513 | 1510 | "setsum f A = msetsum (image_mset f (mset_set A))" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1511 | by (cases "finite A") (induct A rule: finite_induct, simp_all) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1512 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1513 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1514 | |
| 59813 | 1515 | lemma msetsum_diff: | 
| 61076 | 1516 |   fixes M N :: "('a :: ordered_cancel_comm_monoid_diff) multiset"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1517 | shows "N \<subseteq># M \<Longrightarrow> msetsum (M - N) = msetsum M - msetsum N" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 1518 | by (metis add_diff_cancel_right' msetsum.union subset_mset.diff_add) | 
| 59813 | 1519 | |
| 59949 | 1520 | lemma size_eq_msetsum: "size M = msetsum (image_mset (\<lambda>_. 1) M)" | 
| 1521 | proof (induct M) | |
| 1522 | case empty then show ?case by simp | |
| 1523 | next | |
| 1524 | case (add M x) then show ?case | |
| 60495 | 1525 | by (cases "x \<in> set_mset M") | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1526 | (simp_all add: size_multiset_overloaded_eq setsum.distrib setsum.delta' insert_absorb not_in_iff) | 
| 59949 | 1527 | qed | 
| 1528 | ||
| 62366 | 1529 | syntax (ASCII) | 
| 1530 |   "_msetsum_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3SUM _:#_. _)" [0, 51, 10] 10)
 | |
| 1531 | syntax | |
| 1532 |   "_msetsum_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10)
 | |
| 1533 | translations | |
| 1534 | "\<Sum>i \<in># A. b" \<rightleftharpoons> "CONST msetsum (CONST image_mset (\<lambda>i. b) A)" | |
| 59949 | 1535 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1536 | abbreviation Union_mset :: "'a multiset multiset \<Rightarrow> 'a multiset"  ("\<Union>#_" [900] 900)
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1537 | where "\<Union># MM \<equiv> msetsum MM" -- \<open>FIXME ambiguous notation -- | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1538 |     could likewise refer to @{text "\<Squnion>#"}\<close>
 | 
| 59813 | 1539 | |
| 60495 | 1540 | lemma set_mset_Union_mset[simp]: "set_mset (\<Union># MM) = (\<Union>M \<in> set_mset MM. set_mset M)" | 
| 59813 | 1541 | by (induct MM) auto | 
| 1542 | ||
| 1543 | lemma in_Union_mset_iff[iff]: "x \<in># \<Union># MM \<longleftrightarrow> (\<exists>M. M \<in># MM \<and> x \<in># M)" | |
| 1544 | by (induct MM) auto | |
| 1545 | ||
| 62366 | 1546 | lemma count_setsum: | 
| 1547 | "count (setsum f A) x = setsum (\<lambda>a. count (f a) x) A" | |
| 1548 | by (induct A rule: infinite_finite_induct) simp_all | |
| 1549 | ||
| 1550 | lemma setsum_eq_empty_iff: | |
| 1551 | assumes "finite A" | |
| 1552 |   shows "setsum f A = {#} \<longleftrightarrow> (\<forall>a\<in>A. f a = {#})"
 | |
| 1553 | using assms by induct simp_all | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1554 | |
| 54868 | 1555 | context comm_monoid_mult | 
| 1556 | begin | |
| 1557 | ||
| 61605 | 1558 | sublocale msetprod: comm_monoid_mset times 1 | 
| 61832 | 1559 | defines msetprod = msetprod.F .. | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1560 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1561 | lemma msetprod_empty: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1562 |   "msetprod {#} = 1"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1563 | by (fact msetprod.empty) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1564 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1565 | lemma msetprod_singleton: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1566 |   "msetprod {#x#} = x"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1567 | by (fact msetprod.singleton) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1568 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1569 | lemma msetprod_Un: | 
| 58425 | 1570 | "msetprod (A + B) = msetprod A * msetprod B" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1571 | by (fact msetprod.union) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1572 | |
| 60804 | 1573 | lemma msetprod_replicate_mset [simp]: | 
| 1574 | "msetprod (replicate_mset n a) = a ^ n" | |
| 1575 | by (induct n) (simp_all add: ac_simps) | |
| 1576 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1577 | lemma setprod_unfold_msetprod: | 
| 60513 | 1578 | "setprod f A = msetprod (image_mset f (mset_set A))" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1579 | by (cases "finite A") (induct A rule: finite_induct, simp_all) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1580 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1581 | lemma msetprod_multiplicity: | 
| 60495 | 1582 | "msetprod M = setprod (\<lambda>x. x ^ count M x) (set_mset M)" | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1583 | by (simp add: fold_mset_def setprod.eq_fold msetprod.eq_fold funpow_times_power comp_def) | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1584 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1585 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1586 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1587 | syntax (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1588 |   "_msetprod_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3PROD _:#_. _)" [0, 51, 10] 10)
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1589 | syntax | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1590 |   "_msetprod_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3\<Prod>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1591 | translations | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1592 | "\<Prod>i \<in># A. b" \<rightleftharpoons> "CONST msetprod (CONST image_mset (\<lambda>i. b) A)" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1593 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1594 | lemma (in comm_semiring_1) dvd_msetprod: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1595 | assumes "x \<in># A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1596 | shows "x dvd msetprod A" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1597 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1598 |   from assms have "A = (A - {#x#}) + {#x#}" by simp
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1599 |   then obtain B where "A = B + {#x#}" ..
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1600 | then show ?thesis by simp | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1601 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1602 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1603 | lemma (in semidom) msetprod_zero_iff [iff]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1604 | "msetprod A = 0 \<longleftrightarrow> 0 \<in># A" | 
| 62366 | 1605 | by (induct A) auto | 
| 1606 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1607 | lemma (in semidom_divide) msetprod_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1608 | assumes "B \<subseteq># A" and "0 \<notin># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1609 | shows "msetprod (A - B) = msetprod A div msetprod B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1610 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1611 | from assms obtain C where "A = B + C" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1612 | by (metis subset_mset.add_diff_inverse) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1613 | with assms show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1614 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1615 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1616 | lemma (in semidom_divide) msetprod_minus: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1617 | assumes "a \<in># A" and "a \<noteq> 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1618 |   shows "msetprod (A - {#a#}) = msetprod A div a"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1619 |   using assms msetprod_diff [of "{#a#}" A]
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1620 | by (auto simp add: single_subset_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1621 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1622 | lemma (in normalization_semidom) normalized_msetprodI: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1623 | assumes "\<And>a. a \<in># A \<Longrightarrow> normalize a = a" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1624 | shows "normalize (msetprod A) = msetprod A" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1625 | using assms by (induct A) (simp_all add: normalize_mult) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1626 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1627 | |
| 60500 | 1628 | subsection \<open>Alternative representations\<close> | 
| 1629 | ||
| 1630 | subsubsection \<open>Lists\<close> | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 1631 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1632 | context linorder | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1633 | begin | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1634 | |
| 60515 | 1635 | lemma mset_insort [simp]: | 
| 1636 |   "mset (insort_key k x xs) = {#x#} + mset xs"
 | |
| 37107 | 1637 | by (induct xs) (simp_all add: ac_simps) | 
| 58425 | 1638 | |
| 60515 | 1639 | lemma mset_sort [simp]: | 
| 1640 | "mset (sort_key k xs) = mset xs" | |
| 37107 | 1641 | by (induct xs) (simp_all add: ac_simps) | 
| 1642 | ||
| 60500 | 1643 | text \<open> | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1644 | This lemma shows which properties suffice to show that a function | 
| 61585 | 1645 | \<open>f\<close> with \<open>f xs = ys\<close> behaves like sort. | 
| 60500 | 1646 | \<close> | 
| 37074 | 1647 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1648 | lemma properties_for_sort_key: | 
| 60515 | 1649 | assumes "mset ys = mset xs" | 
| 60606 | 1650 | and "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>x. f k = f x) ys = filter (\<lambda>x. f k = f x) xs" | 
| 1651 | and "sorted (map f ys)" | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1652 | shows "sort_key f xs = ys" | 
| 60606 | 1653 | using assms | 
| 46921 | 1654 | 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 | 1655 | 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 | 1656 | 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 | 1657 | case (Cons x xs) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1658 | from Cons.prems(2) have | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1659 | "\<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 | 1660 | by (simp add: filter_remove1) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1661 | 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 | 1662 | by (auto intro!: Cons.hyps simp add: sorted_map_remove1) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1663 | moreover from Cons.prems have "x \<in># mset ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1664 | by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1665 | then have "x \<in> set ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1666 | by simp | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1667 | 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 | 1668 | 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 | 1669 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1670 | lemma properties_for_sort: | 
| 60515 | 1671 | assumes multiset: "mset ys = mset xs" | 
| 60606 | 1672 | and "sorted ys" | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1673 | shows "sort xs = ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1674 | proof (rule properties_for_sort_key) | 
| 60515 | 1675 | from multiset show "mset ys = mset xs" . | 
| 60500 | 1676 | from \<open>sorted ys\<close> show "sorted (map (\<lambda>x. x) ys)" by simp | 
| 60678 | 1677 | from multiset have "length (filter (\<lambda>y. k = y) ys) = length (filter (\<lambda>x. k = x) xs)" for k | 
| 60515 | 1678 | by (rule mset_eq_length_filter) | 
| 60678 | 1679 | then have "replicate (length (filter (\<lambda>y. k = y) ys)) k = | 
| 1680 | replicate (length (filter (\<lambda>x. k = x) xs)) k" for k | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1681 | by simp | 
| 60678 | 1682 | then show "k \<in> set ys \<Longrightarrow> filter (\<lambda>y. k = y) ys = filter (\<lambda>x. k = x) xs" for k | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1683 | 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 | 1684 | qed | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1685 | |
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1686 | lemma sort_key_inj_key_eq: | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1687 | assumes mset_equal: "mset xs = mset ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1688 | and "inj_on f (set xs)" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1689 | and "sorted (map f ys)" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1690 | shows "sort_key f xs = ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1691 | proof (rule properties_for_sort_key) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1692 | from mset_equal | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1693 | show "mset ys = mset xs" by simp | 
| 61188 | 1694 | from \<open>sorted (map f ys)\<close> | 
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1695 | show "sorted (map f ys)" . | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1696 | show "[x\<leftarrow>ys . f k = f x] = [x\<leftarrow>xs . f k = f x]" if "k \<in> set ys" for k | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1697 | proof - | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1698 | from mset_equal | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1699 | have set_equal: "set xs = set ys" by (rule mset_eq_setD) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1700 | with that have "insert k (set ys) = set ys" by auto | 
| 61188 | 1701 | with \<open>inj_on f (set xs)\<close> have inj: "inj_on f (insert k (set ys))" | 
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1702 | by (simp add: set_equal) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1703 | from inj have "[x\<leftarrow>ys . f k = f x] = filter (HOL.eq k) ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1704 | by (auto intro!: inj_on_filter_key_eq) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1705 | also have "\<dots> = replicate (count (mset ys) k) k" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1706 | by (simp add: replicate_count_mset_eq_filter_eq) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1707 | also have "\<dots> = replicate (count (mset xs) k) k" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1708 | using mset_equal by simp | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1709 | also have "\<dots> = filter (HOL.eq k) xs" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1710 | by (simp add: replicate_count_mset_eq_filter_eq) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1711 | also have "\<dots> = [x\<leftarrow>xs . f k = f x]" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1712 | using inj by (auto intro!: inj_on_filter_key_eq [symmetric] simp add: set_equal) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1713 | finally show ?thesis . | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1714 | qed | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1715 | qed | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1716 | |
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1717 | lemma sort_key_eq_sort_key: | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1718 | assumes "mset xs = mset ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1719 | and "inj_on f (set xs)" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1720 | shows "sort_key f xs = sort_key f ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1721 | by (rule sort_key_inj_key_eq) (simp_all add: assms) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 1722 | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1723 | lemma sort_key_by_quicksort: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1724 | "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 | 1725 | @ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1726 | @ 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 | 1727 | proof (rule properties_for_sort_key) | 
| 60515 | 1728 | show "mset ?rhs = mset ?lhs" | 
| 1729 | by (rule multiset_eqI) (auto simp add: mset_filter) | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1730 | show "sorted (map f ?rhs)" | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1731 | by (auto simp add: sorted_append intro: sorted_map_same) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1732 | next | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1733 | fix l | 
| 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1734 | assume "l \<in> set ?rhs" | 
| 40346 | 1735 | let ?pivot = "f (xs ! (length xs div 2))" | 
| 1736 | have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto | |
| 40306 | 1737 | 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 | 1738 | unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) | 
| 40346 | 1739 | with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp | 
| 1740 | 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 | |
| 1741 | then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] = | |
| 1742 | [x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp | |
| 1743 | note *** = this [of "op <"] this [of "op >"] this [of "op ="] | |
| 40306 | 1744 | 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 | 1745 | proof (cases "f l" ?pivot rule: linorder_cases) | 
| 46730 | 1746 | case less | 
| 1747 | then have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto | |
| 1748 | with less show ?thesis | |
| 40346 | 1749 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1750 | next | 
| 40306 | 1751 | case equal then show ?thesis | 
| 40346 | 1752 | by (simp add: * less_le) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 1753 | next | 
| 46730 | 1754 | case greater | 
| 1755 | then have "f l \<noteq> ?pivot" and "\<not> f l < ?pivot" by auto | |
| 1756 | with greater show ?thesis | |
| 40346 | 1757 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40306 | 1758 | qed | 
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1759 | qed | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1760 | |
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1761 | lemma sort_by_quicksort: | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1762 | "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 | 1763 | @ [x\<leftarrow>xs. x = xs ! (length xs div 2)] | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 1764 | @ 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 | 1765 | 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 | 1766 | |
| 60500 | 1767 | text \<open>A stable parametrized quicksort\<close> | 
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1768 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1769 | 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 | 1770 | "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 | 1771 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1772 | lemma part_code [code]: | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1773 | "part f pivot [] = ([], [], [])" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1774 | "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 | 1775 | if x' < pivot then (x # lts, eqs, gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1776 | else if x' > pivot then (lts, eqs, x # gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1777 | else (lts, x # eqs, gts))" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1778 | by (auto simp add: part_def Let_def split_def) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1779 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1780 | lemma sort_key_by_quicksort_code [code]: | 
| 60606 | 1781 | "sort_key f xs = | 
| 1782 | (case xs of | |
| 1783 | [] \<Rightarrow> [] | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1784 | | [x] \<Rightarrow> xs | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1785 | | [x, y] \<Rightarrow> (if f x \<le> f y then xs else [y, x]) | 
| 60606 | 1786 | | _ \<Rightarrow> | 
| 1787 | let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs | |
| 1788 | in sort_key f lts @ eqs @ sort_key f gts)" | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1789 | proof (cases xs) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1790 | case Nil then show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1791 | next | 
| 46921 | 1792 | case (Cons _ ys) note hyps = Cons show ?thesis | 
| 1793 | proof (cases ys) | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1794 | case Nil with hyps show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1795 | next | 
| 46921 | 1796 | case (Cons _ zs) note hyps = hyps Cons show ?thesis | 
| 1797 | proof (cases zs) | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1798 | case Nil with hyps show ?thesis by auto | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1799 | next | 
| 58425 | 1800 | case Cons | 
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1801 | from sort_key_by_quicksort [of f xs] | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1802 | 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 | 1803 | in sort_key f lts @ eqs @ sort_key f gts)" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1804 | 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 | 1805 | with hyps Cons show ?thesis by (simp only: list.cases) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1806 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1807 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1808 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1809 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1810 | end | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 1811 | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1812 | hide_const (open) part | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 1813 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1814 | lemma mset_remdups_le: "mset (remdups xs) \<subseteq># mset xs" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 1815 | by (induct xs) (auto intro: subset_mset.order_trans) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1816 | |
| 60515 | 1817 | lemma mset_update: | 
| 1818 |   "i < length ls \<Longrightarrow> mset (ls[i := v]) = mset ls - {#ls ! i#} + {#v#}"
 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1819 | 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 | 1820 | 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 | 1821 | 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 | 1822 | 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 | 1823 | 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 | 1824 | 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 | 1825 | 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 | 1826 | 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 | 1827 | 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 | 1828 | 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 | 1829 | apply simp | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1830 | apply (subst add.assoc) | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1831 |       apply (subst add.commute [of "{#v#}" "{#x#}"])
 | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1832 | 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 | 1833 | 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 | 1834 | apply (rule mset_le_multiset_union_diff_commute) | 
| 60515 | 1835 | apply (simp add: mset_le_single nth_mem_mset) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1836 | 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 | 1837 | 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 | 1838 | 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 | 1839 | |
| 60515 | 1840 | lemma mset_swap: | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1841 | "i < length ls \<Longrightarrow> j < length ls \<Longrightarrow> | 
| 60515 | 1842 | mset (ls[j := ls ! i, i := ls ! j]) = mset ls" | 
| 1843 | by (cases "i = j") (simp_all add: mset_update nth_mem_mset) | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1844 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1845 | |
| 60500 | 1846 | subsection \<open>The multiset order\<close> | 
| 1847 | ||
| 1848 | subsubsection \<open>Well-foundedness\<close> | |
| 10249 | 1849 | |
| 60606 | 1850 | definition mult1 :: "('a \<times> 'a) set \<Rightarrow> ('a multiset \<times> 'a multiset) set" where
 | 
| 37765 | 1851 |   "mult1 r = {(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
 | 
| 60607 | 1852 | (\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r)}" | 
| 60606 | 1853 | |
| 1854 | definition mult :: "('a \<times> 'a) set \<Rightarrow> ('a multiset \<times> 'a multiset) set" where
 | |
| 37765 | 1855 | "mult r = (mult1 r)\<^sup>+" | 
| 10249 | 1856 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1857 | lemma mult1I: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1858 |   assumes "M = M0 + {#a#}" and "N = M0 + K" and "\<And>b. b \<in># K \<Longrightarrow> (b, a) \<in> r"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1859 | shows "(N, M) \<in> mult1 r" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1860 | using assms unfolding mult1_def by blast | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1861 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1862 | lemma mult1E: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1863 | assumes "(N, M) \<in> mult1 r" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1864 |   obtains a M0 K where "M = M0 + {#a#}" "N = M0 + K" "\<And>b. b \<in># K \<Longrightarrow> (b, a) \<in> r"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1865 | using assms unfolding mult1_def by blast | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1866 | |
| 23751 | 1867 | lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
 | 
| 26178 | 1868 | by (simp add: mult1_def) | 
| 10249 | 1869 | |
| 60608 | 1870 | lemma less_add: | 
| 1871 |   assumes mult1: "(N, M0 + {#a#}) \<in> mult1 r"
 | |
| 1872 | shows | |
| 1873 |     "(\<exists>M. (M, M0) \<in> mult1 r \<and> N = M + {#a#}) \<or>
 | |
| 1874 | (\<exists>K. (\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r) \<and> N = M0 + K)" | |
| 1875 | proof - | |
| 60607 | 1876 | let ?r = "\<lambda>K a. \<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r" | 
| 11464 | 1877 |   let ?R = "\<lambda>N M. \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and> ?r K a"
 | 
| 60608 | 1878 |   obtain a' M0' K where M0: "M0 + {#a#} = M0' + {#a'#}"
 | 
| 1879 | and N: "N = M0' + K" | |
| 1880 | and r: "?r K a'" | |
| 1881 | using mult1 unfolding mult1_def by auto | |
| 1882 | show ?thesis (is "?case1 \<or> ?case2") | |
| 60606 | 1883 | proof - | 
| 1884 | from M0 consider "M0 = M0'" "a = a'" | |
| 1885 |       | K' where "M0 = K' + {#a'#}" "M0' = K' + {#a#}"
 | |
| 1886 | by atomize_elim (simp only: add_eq_conv_ex) | |
| 18258 | 1887 | then show ?thesis | 
| 60606 | 1888 | proof cases | 
| 1889 | case 1 | |
| 11464 | 1890 | with N r have "?r K a \<and> N = M0 + K" by simp | 
| 60606 | 1891 | then have ?case2 .. | 
| 1892 | then show ?thesis .. | |
| 10249 | 1893 | next | 
| 60606 | 1894 | case 2 | 
| 1895 |       from N 2(2) have n: "N = K' + K + {#a#}" by (simp add: ac_simps)
 | |
| 1896 | with r 2(1) have "?R (K' + K) M0" by blast | |
| 60608 | 1897 | with n have ?case1 by (simp add: mult1_def) | 
| 60606 | 1898 | then show ?thesis .. | 
| 10249 | 1899 | qed | 
| 1900 | qed | |
| 1901 | qed | |
| 1902 | ||
| 60608 | 1903 | lemma all_accessible: | 
| 1904 | assumes "wf r" | |
| 1905 | shows "\<forall>M. M \<in> Wellfounded.acc (mult1 r)" | |
| 10249 | 1906 | proof | 
| 1907 | let ?R = "mult1 r" | |
| 54295 | 1908 | let ?W = "Wellfounded.acc ?R" | 
| 10249 | 1909 |   {
 | 
| 1910 | fix M M0 a | |
| 23751 | 1911 | assume M0: "M0 \<in> ?W" | 
| 60606 | 1912 |       and wf_hyp: "\<And>b. (b, a) \<in> r \<Longrightarrow> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | 
| 1913 |       and acc_hyp: "\<forall>M. (M, M0) \<in> ?R \<longrightarrow> M + {#a#} \<in> ?W"
 | |
| 23751 | 1914 |     have "M0 + {#a#} \<in> ?W"
 | 
| 1915 |     proof (rule accI [of "M0 + {#a#}"])
 | |
| 10249 | 1916 | fix N | 
| 23751 | 1917 |       assume "(N, M0 + {#a#}) \<in> ?R"
 | 
| 60608 | 1918 |       then consider M where "(M, M0) \<in> ?R" "N = M + {#a#}"
 | 
| 1919 | | K where "\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r" "N = M0 + K" | |
| 1920 | by atomize_elim (rule less_add) | |
| 23751 | 1921 | then show "N \<in> ?W" | 
| 60608 | 1922 | proof cases | 
| 1923 | case 1 | |
| 60606 | 1924 |         from acc_hyp have "(M, M0) \<in> ?R \<longrightarrow> M + {#a#} \<in> ?W" ..
 | 
| 60500 | 1925 |         from this and \<open>(M, M0) \<in> ?R\<close> have "M + {#a#} \<in> ?W" ..
 | 
| 60608 | 1926 |         then show "N \<in> ?W" by (simp only: \<open>N = M + {#a#}\<close>)
 | 
| 10249 | 1927 | next | 
| 60608 | 1928 | case 2 | 
| 1929 | from this(1) have "M0 + K \<in> ?W" | |
| 10249 | 1930 | proof (induct K) | 
| 18730 | 1931 | case empty | 
| 23751 | 1932 |           from M0 show "M0 + {#} \<in> ?W" by simp
 | 
| 18730 | 1933 | next | 
| 1934 | case (add K x) | |
| 23751 | 1935 | from add.prems have "(x, a) \<in> r" by simp | 
| 1936 |           with wf_hyp have "\<forall>M \<in> ?W. M + {#x#} \<in> ?W" by blast
 | |
| 1937 | moreover from add have "M0 + K \<in> ?W" by simp | |
| 1938 |           ultimately have "(M0 + K) + {#x#} \<in> ?W" ..
 | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 1939 |           then show "M0 + (K + {#x#}) \<in> ?W" by (simp only: add.assoc)
 | 
| 10249 | 1940 | qed | 
| 60608 | 1941 | then show "N \<in> ?W" by (simp only: 2(2)) | 
| 10249 | 1942 | qed | 
| 1943 | qed | |
| 1944 | } note tedious_reasoning = this | |
| 1945 | ||
| 60608 | 1946 | show "M \<in> ?W" for M | 
| 10249 | 1947 | proof (induct M) | 
| 23751 | 1948 |     show "{#} \<in> ?W"
 | 
| 10249 | 1949 | proof (rule accI) | 
| 23751 | 1950 |       fix b assume "(b, {#}) \<in> ?R"
 | 
| 1951 | with not_less_empty show "b \<in> ?W" by contradiction | |
| 10249 | 1952 | qed | 
| 1953 | ||
| 23751 | 1954 | fix M a assume "M \<in> ?W" | 
| 60608 | 1955 |     from \<open>wf r\<close> have "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | 
| 10249 | 1956 | proof induct | 
| 1957 | fix a | |
| 60606 | 1958 |       assume r: "\<And>b. (b, a) \<in> r \<Longrightarrow> (\<forall>M \<in> ?W. M + {#b#} \<in> ?W)"
 | 
| 23751 | 1959 |       show "\<forall>M \<in> ?W. M + {#a#} \<in> ?W"
 | 
| 10249 | 1960 | proof | 
| 23751 | 1961 | fix M assume "M \<in> ?W" | 
| 1962 |         then show "M + {#a#} \<in> ?W"
 | |
| 23373 | 1963 | by (rule acc_induct) (rule tedious_reasoning [OF _ r]) | 
| 10249 | 1964 | qed | 
| 1965 | qed | |
| 60500 | 1966 |     from this and \<open>M \<in> ?W\<close> show "M + {#a#} \<in> ?W" ..
 | 
| 10249 | 1967 | qed | 
| 1968 | qed | |
| 1969 | ||
| 60606 | 1970 | theorem wf_mult1: "wf r \<Longrightarrow> wf (mult1 r)" | 
| 26178 | 1971 | by (rule acc_wfI) (rule all_accessible) | 
| 10249 | 1972 | |
| 60606 | 1973 | theorem wf_mult: "wf r \<Longrightarrow> wf (mult r)" | 
| 26178 | 1974 | unfolding mult_def by (rule wf_trancl) (rule wf_mult1) | 
| 10249 | 1975 | |
| 1976 | ||
| 60500 | 1977 | subsubsection \<open>Closure-free presentation\<close> | 
| 1978 | ||
| 1979 | text \<open>One direction.\<close> | |
| 10249 | 1980 | |
| 1981 | lemma mult_implies_one_step: | |
| 60606 | 1982 | "trans r \<Longrightarrow> (M, N) \<in> mult r \<Longrightarrow> | 
| 11464 | 1983 |     \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and>
 | 
| 60495 | 1984 | (\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> r)" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1985 | apply (unfold mult_def mult1_def) | 
| 26178 | 1986 | apply (erule converse_trancl_induct, clarify) | 
| 1987 | apply (rule_tac x = M0 in exI, simp, clarify) | |
| 60607 | 1988 | apply (case_tac "a \<in># K") | 
| 26178 | 1989 | apply (rule_tac x = I in exI) | 
| 1990 | apply (simp (no_asm)) | |
| 1991 |  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 | 1992 | apply (simp (no_asm_simp) add: add.assoc [symmetric]) | 
| 59807 | 1993 |  apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="S + T" for S T in arg_cong)
 | 
| 26178 | 1994 | apply (simp add: diff_union_single_conv) | 
| 1995 | apply (simp (no_asm_use) add: trans_def) | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1996 | apply (metis (no_types, hide_lams) Multiset.diff_right_commute Un_iff diff_single_trivial multi_drop_mem_not_eq) | 
| 60607 | 1997 | apply (subgoal_tac "a \<in># I") | 
| 26178 | 1998 |  apply (rule_tac x = "I - {#a#}" in exI)
 | 
| 1999 |  apply (rule_tac x = "J + {#a#}" in exI)
 | |
| 2000 | apply (rule_tac x = "K + Ka" in exI) | |
| 2001 | 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 | 2002 | apply (simp add: multiset_eq_iff split: nat_diff_split) | 
| 26178 | 2003 | apply (rule conjI) | 
| 59807 | 2004 |   apply (drule_tac f = "\<lambda>M. M - {#a#}" and x="S + T" for 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 | 2005 | apply (simp add: multiset_eq_iff split: nat_diff_split) | 
| 26178 | 2006 | apply (simp (no_asm_use) add: trans_def) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2007 | apply (subgoal_tac "a \<in># (M0 + {#a#})")
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2008 | apply (simp_all add: not_in_iff) | 
| 26178 | 2009 | apply blast | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2010 | apply (metis add.comm_neutral add_diff_cancel_right' count_eq_zero_iff diff_single_trivial multi_self_add_other_not_self plus_multiset.rep_eq) | 
| 26178 | 2011 | done | 
| 10249 | 2012 | |
| 2013 | lemma one_step_implies_mult_aux: | |
| 60678 | 2014 |   "\<forall>I J K. size J = n \<and> J \<noteq> {#} \<and> (\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> r)
 | 
| 2015 | \<longrightarrow> (I + K, I + J) \<in> mult r" | |
| 2016 | apply (induct n) | |
| 2017 | apply auto | |
| 26178 | 2018 | apply (frule size_eq_Suc_imp_eq_union, clarify) | 
| 2019 | apply (rename_tac "J'", simp) | |
| 2020 | apply (erule notE, auto) | |
| 2021 | apply (case_tac "J' = {#}")
 | |
| 2022 | apply (simp add: mult_def) | |
| 2023 | apply (rule r_into_trancl) | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2024 | apply (simp add: mult1_def, blast) | 
| 60500 | 2025 | txt \<open>Now we know @{term "J' \<noteq> {#}"}.\<close>
 | 
| 26178 | 2026 | apply (cut_tac M = K and P = "\<lambda>x. (x, a) \<in> r" in multiset_partition) | 
| 60495 | 2027 | apply (erule_tac P = "\<forall>k \<in> set_mset K. P k" for P in rev_mp) | 
| 26178 | 2028 | apply (erule ssubst) | 
| 2029 | apply (simp add: Ball_def, auto) | |
| 2030 | apply (subgoal_tac | |
| 60607 | 2031 |   "((I + {# x \<in># K. (x, a) \<in> r #}) + {# x \<in># K. (x, a) \<notin> r #},
 | 
| 2032 |     (I + {# x \<in># K. (x, a) \<in> r #}) + J') \<in> mult r")
 | |
| 26178 | 2033 | prefer 2 | 
| 2034 | apply force | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 2035 | apply (simp (no_asm_use) add: add.assoc [symmetric] mult_def) | 
| 26178 | 2036 | apply (erule trancl_trans) | 
| 2037 | apply (rule r_into_trancl) | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2038 | apply (simp add: mult1_def) | 
| 26178 | 2039 | apply (rule_tac x = a in exI) | 
| 2040 | 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 | 2041 | apply (simp add: ac_simps) | 
| 26178 | 2042 | done | 
| 10249 | 2043 | |
| 17161 | 2044 | lemma one_step_implies_mult: | 
| 60606 | 2045 |   "trans r \<Longrightarrow> J \<noteq> {#} \<Longrightarrow> \<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> r
 | 
| 2046 | \<Longrightarrow> (I + K, I + J) \<in> mult r" | |
| 26178 | 2047 | using one_step_implies_mult_aux by blast | 
| 10249 | 2048 | |
| 2049 | ||
| 60500 | 2050 | subsubsection \<open>Partial-order properties\<close> | 
| 10249 | 2051 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2052 | lemma (in order) mult1_lessE: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2053 |   assumes "(N, M) \<in> mult1 {(a, b). a < b}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2054 |   obtains a M0 K where "M = M0 + {#a#}" "N = M0 + K"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2055 | "a \<notin># K" "\<And>b. b \<in># K \<Longrightarrow> b < a" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2056 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2057 |   from assms obtain a M0 K where "M = M0 + {#a#}" "N = M0 + K"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2058 | "\<And>b. b \<in># K \<Longrightarrow> b < a" by (blast elim: mult1E) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2059 | moreover from this(3) [of a] have "a \<notin># K" by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2060 | ultimately show thesis by (auto intro: that) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2061 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2062 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2063 | definition less_multiset :: "'a::order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "#\<subset>#" 50) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2064 |   where "M' #\<subset># M \<longleftrightarrow> (M', M) \<in> mult {(x', x). x' < x}"
 | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2065 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2066 | definition le_multiset :: "'a::order multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "#\<subseteq>#" 50) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2067 | where "M' #\<subseteq># M \<longleftrightarrow> M' #\<subset># M \<or> M' = M" | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2068 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2069 | notation (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2070 | less_multiset (infix "#<#" 50) and | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2071 | le_multiset (infix "#<=#" 50) | 
| 10249 | 2072 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2073 | interpretation multiset_order: order le_multiset less_multiset | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2074 | proof - | 
| 60606 | 2075 | have irrefl: "\<not> M #\<subset># M" for M :: "'a multiset" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2076 | proof | 
| 59958 
4538d41e8e54
renamed multiset ordering to free up nice <# etc. symbols for the standard subset
 blanchet parents: 
59949diff
changeset | 2077 | assume "M #\<subset># M" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2078 |     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 | 2079 |     have "trans {(x'::'a, x). x' < x}"
 | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2080 | by (rule transI) simp | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2081 | moreover note MM | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2082 | ultimately have "\<exists>I J K. M = I + J \<and> M = I + K | 
| 60495 | 2083 |       \<and> J \<noteq> {#} \<and> (\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset J. (k, j) \<in> {(x, y). x < y})"
 | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2084 | by (rule mult_implies_one_step) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2085 | then obtain I J K where "M = I + J" and "M = I + K" | 
| 60495 | 2086 |       and "J \<noteq> {#}" and "(\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset J. (k, j) \<in> {(x, y). x < y})" by blast
 | 
| 60678 | 2087 |     then have *: "K \<noteq> {#}" and **: "\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset K. k < j" by auto
 | 
| 60495 | 2088 | have "finite (set_mset K)" by simp | 
| 60678 | 2089 | moreover note ** | 
| 60495 | 2090 |     ultimately have "set_mset K = {}"
 | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2091 | by (induct rule: finite_induct) (auto intro: order_less_trans) | 
| 60678 | 2092 | with * show False by simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2093 | qed | 
| 60678 | 2094 | have trans: "K #\<subset># M \<Longrightarrow> M #\<subset># N \<Longrightarrow> K #\<subset># N" for K M N :: "'a multiset" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2095 | unfolding less_multiset_def mult_def by (blast intro: trancl_trans) | 
| 46921 | 2096 | show "class.order (le_multiset :: 'a multiset \<Rightarrow> _) less_multiset" | 
| 60678 | 2097 | by standard (auto simp add: le_multiset_def irrefl dest: trans) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2098 | qed -- \<open>FIXME avoid junk stemming from type class interpretation\<close> | 
| 10249 | 2099 | |
| 60678 | 2100 | lemma mult_less_irrefl [elim!]: | 
| 2101 | fixes M :: "'a::order multiset" | |
| 2102 | shows "M #\<subset># M \<Longrightarrow> R" | |
| 46730 | 2103 | by simp | 
| 26567 
7bcebb8c2d33
instantiation replacing primitive instance plus overloaded defs; more conservative type arities
 haftmann parents: 
26178diff
changeset | 2104 | |
| 10249 | 2105 | |
| 60500 | 2106 | subsubsection \<open>Monotonicity of multiset union\<close> | 
| 10249 | 2107 | |
| 60606 | 2108 | lemma mult1_union: "(B, D) \<in> mult1 r \<Longrightarrow> (C + B, C + D) \<in> mult1 r" | 
| 26178 | 2109 | apply (unfold mult1_def) | 
| 2110 | apply auto | |
| 2111 | apply (rule_tac x = a in exI) | |
| 2112 | 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 | 2113 | apply (simp add: add.assoc) | 
| 26178 | 2114 | done | 
| 10249 | 2115 | |
| 60606 | 2116 | lemma union_less_mono2: "B #\<subset># D \<Longrightarrow> C + B #\<subset># C + (D::'a::order multiset)" | 
| 26178 | 2117 | apply (unfold less_multiset_def mult_def) | 
| 2118 | apply (erule trancl_induct) | |
| 40249 
cd404ecb9107
Remove unnecessary premise of mult1_union
 Lars Noschinski <noschinl@in.tum.de> parents: 
39533diff
changeset | 2119 | apply (blast intro: mult1_union) | 
| 
cd404ecb9107
Remove unnecessary premise of mult1_union
 Lars Noschinski <noschinl@in.tum.de> parents: 
39533diff
changeset | 2120 | apply (blast intro: mult1_union trancl_trans) | 
| 26178 | 2121 | done | 
| 10249 | 2122 | |
| 60606 | 2123 | lemma union_less_mono1: "B #\<subset># D \<Longrightarrow> 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 | 2124 | apply (subst add.commute [of B C]) | 
| 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57492diff
changeset | 2125 | apply (subst add.commute [of D C]) | 
| 26178 | 2126 | apply (erule union_less_mono2) | 
| 2127 | done | |
| 10249 | 2128 | |
| 17161 | 2129 | lemma union_less_mono: | 
| 60606 | 2130 | fixes A B C D :: "'a::order multiset" | 
| 2131 | shows "A #\<subset># C \<Longrightarrow> B #\<subset># D \<Longrightarrow> A + B #\<subset># C + D" | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2132 | by (blast intro!: union_less_mono1 union_less_mono2 multiset_order.less_trans) | 
| 10249 | 2133 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2134 | interpretation multiset_order: ordered_ab_semigroup_add plus le_multiset less_multiset | 
| 60678 | 2135 | by standard (auto simp add: le_multiset_def intro: union_less_mono2) | 
| 26145 | 2136 | |
| 15072 | 2137 | |
| 60500 | 2138 | subsubsection \<open>Termination proofs with multiset orders\<close> | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2139 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2140 | 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 | 2141 |   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 | 2142 |   and multi_member_last: "x \<in># {# x #}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2143 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2144 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2145 | definition "ms_strict = mult pair_less" | 
| 37765 | 2146 | 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 | 2147 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2148 | 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 | 2149 | 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 | 2150 | 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 | 2151 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2152 | lemma smsI: | 
| 60495 | 2153 | "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z + B) \<in> ms_strict" | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2154 | unfolding ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2155 | 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 | 2156 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2157 | lemma wmsI: | 
| 60495 | 2158 |   "(set_mset A, set_mset B) \<in> max_strict \<or> A = {#} \<and> B = {#}
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2159 | \<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 | 2160 | unfolding ms_weak_def ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2161 | 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 | 2162 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2163 | inductive pw_leq | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2164 | where | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2165 |   pw_leq_empty: "pw_leq {#} {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2166 | | 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 | 2167 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2168 | lemma pw_leq_lstep: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2169 |   "(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 | 2170 | 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 | 2171 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2172 | lemma pw_leq_split: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2173 | assumes "pw_leq X Y" | 
| 60495 | 2174 |   shows "\<exists>A B Z. X = A + Z \<and> Y = B + Z \<and> ((set_mset A, set_mset B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2175 | using assms | 
| 60606 | 2176 | proof induct | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2177 | 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 | 2178 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2179 | 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 | 2180 | then obtain A B Z where | 
| 58425 | 2181 | [simp]: "X = A + Z" "Y = B + Z" | 
| 60495 | 2182 |       and 1[simp]: "(set_mset A, set_mset B) \<in> max_strict \<or> (B = {#} \<and> A = {#})"
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2183 | by auto | 
| 60606 | 2184 | from pw_leq_step consider "x = y" | "(x, y) \<in> pair_less" | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2185 | unfolding pair_leq_def by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2186 | thus ?case | 
| 60606 | 2187 | proof cases | 
| 2188 | case [simp]: 1 | |
| 2189 |     have "{#x#} + X = A + ({#y#}+Z) \<and> {#y#} + Y = B + ({#y#}+Z) \<and>
 | |
| 2190 |       ((set_mset A, set_mset 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 | 2191 | by (auto simp: ac_simps) | 
| 60606 | 2192 | thus ?thesis by blast | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2193 | next | 
| 60606 | 2194 | case 2 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2195 |     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 | 2196 |     have "{#x#} + X = ?A' + Z"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2197 |       "{#y#} + Y = ?B' + Z"
 | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
57512diff
changeset | 2198 | by (auto simp add: ac_simps) | 
| 58425 | 2199 | moreover have | 
| 60495 | 2200 | "(set_mset ?A', set_mset ?B') \<in> max_strict" | 
| 60606 | 2201 | using 1 2 unfolding max_strict_def | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2202 | by (auto elim!: max_ext.cases) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2203 | ultimately show ?thesis by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2204 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2205 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2206 | |
| 58425 | 2207 | lemma | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2208 | assumes pwleq: "pw_leq Z Z'" | 
| 60495 | 2209 | shows ms_strictI: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_strict" | 
| 60606 | 2210 | and ms_weakI1: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_weak" | 
| 2211 |     and ms_weakI2:  "(Z + {#}, Z' + {#}) \<in> ms_weak"
 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2212 | proof - | 
| 58425 | 2213 | from pw_leq_split[OF pwleq] | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2214 | obtain A' B' Z'' | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2215 | where [simp]: "Z = A' + Z''" "Z' = B' + Z''" | 
| 60495 | 2216 |     and mx_or_empty: "(set_mset A', set_mset B') \<in> max_strict \<or> (A' = {#} \<and> B' = {#})"
 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2217 | by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2218 |   {
 | 
| 60495 | 2219 | assume max: "(set_mset A, set_mset B) \<in> max_strict" | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2220 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2221 | 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 | 2222 | proof | 
| 60495 | 2223 | assume max': "(set_mset A', set_mset B') \<in> max_strict" | 
| 2224 | with max have "(set_mset (A + A'), set_mset (B + B')) \<in> max_strict" | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2225 | by (auto simp: max_strict_def intro: max_ext_additive) | 
| 58425 | 2226 | thus ?thesis by (rule smsI) | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2227 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2228 |       assume [simp]: "A' = {#} \<and> B' = {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2229 | 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 | 2230 | qed | 
| 60606 | 2231 | 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 | 2232 | 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 | 2233 | } | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2234 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2235 | 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 | 2236 |   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 | 2237 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2238 | |
| 39301 | 2239 | 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 | 2240 | and nonempty_plus: "{# x #} + rs \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2241 | and nonempty_single: "{# x #} \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2242 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2243 | |
| 60500 | 2244 | setup \<open> | 
| 60606 | 2245 | let | 
| 2246 |     fun msetT T = Type (@{type_name multiset}, [T]);
 | |
| 2247 | ||
| 2248 |     fun mk_mset T [] = Const (@{const_abbrev Mempty}, msetT T)
 | |
| 2249 |       | mk_mset T [x] = Const (@{const_name single}, T --> msetT T) $ x
 | |
| 2250 | | mk_mset T (x :: xs) = | |
| 2251 |             Const (@{const_name plus}, msetT T --> msetT T --> msetT T) $
 | |
| 2252 | mk_mset T [x] $ mk_mset T xs | |
| 2253 | ||
| 60752 | 2254 | fun mset_member_tac ctxt m i = | 
| 60606 | 2255 | if m <= 0 then | 
| 60752 | 2256 |         resolve_tac ctxt @{thms multi_member_this} i ORELSE
 | 
| 2257 |         resolve_tac ctxt @{thms multi_member_last} i
 | |
| 60606 | 2258 | else | 
| 60752 | 2259 |         resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i
 | 
| 2260 | ||
| 2261 | fun mset_nonempty_tac ctxt = | |
| 2262 |       resolve_tac ctxt @{thms nonempty_plus} ORELSE'
 | |
| 2263 |       resolve_tac ctxt @{thms nonempty_single}
 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 2264 | |
| 60606 | 2265 | fun regroup_munion_conv ctxt = | 
| 2266 |       Function_Lib.regroup_conv ctxt @{const_abbrev Mempty} @{const_name plus}
 | |
| 2267 |         (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral}))
 | |
| 2268 | ||
| 60752 | 2269 | fun unfold_pwleq_tac ctxt i = | 
| 2270 |       (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st))
 | |
| 2271 |         ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i)
 | |
| 2272 |         ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i)
 | |
| 60606 | 2273 | |
| 2274 |     val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union},
 | |
| 2275 |                         @{thm Un_insert_left}, @{thm Un_empty_left}]
 | |
| 2276 | in | |
| 2277 | ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset | |
| 2278 |     {
 | |
| 2279 | msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, | |
| 2280 | mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, | |
| 2281 | mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps, | |
| 2282 |       smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1},
 | |
| 60752 | 2283 |       reduction_pair = @{thm ms_reduction_pair}
 | 
| 60606 | 2284 | }) | 
| 2285 | end | |
| 60500 | 2286 | \<close> | 
| 2287 | ||
| 2288 | ||
| 2289 | subsection \<open>Legacy theorem bindings\<close> | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2290 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 2291 | 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 | 2292 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2293 | 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 | 2294 | 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 | 2295 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2296 | 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 | 2297 | 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 | 2298 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2299 | 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 | 2300 | 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 | 2301 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2302 | 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 | 2303 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2304 | 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 | 2305 | 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 | 2306 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2307 | 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 | 2308 | 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 | 2309 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2310 | lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y" | 
| 59557 | 2311 | 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 | 2312 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2313 | lemma mset_less_trans: "(M::'a multiset) \<subset># K \<Longrightarrow> K \<subset># N \<Longrightarrow> M \<subset># N" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 2314 | by (fact subset_mset.less_trans) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2315 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2316 | lemma multiset_inter_commute: "A #\<inter> B = B #\<inter> A" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 2317 | by (fact subset_mset.inf.commute) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2318 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2319 | lemma multiset_inter_assoc: "A #\<inter> (B #\<inter> C) = A #\<inter> B #\<inter> C" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 2320 | by (fact subset_mset.inf.assoc [symmetric]) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2321 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2322 | lemma multiset_inter_left_commute: "A #\<inter> (B #\<inter> C) = B #\<inter> (A #\<inter> C)" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 2323 | by (fact subset_mset.inf.left_commute) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2324 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2325 | lemmas multiset_inter_ac = | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2326 | multiset_inter_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2327 | multiset_inter_assoc | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2328 | multiset_inter_left_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2329 | |
| 60606 | 2330 | lemma mult_less_not_refl: "\<not> M #\<subset># (M::'a::order multiset)" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2331 | by (fact multiset_order.less_irrefl) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2332 | |
| 60606 | 2333 | lemma mult_less_trans: "K #\<subset># M \<Longrightarrow> M #\<subset># N \<Longrightarrow> K #\<subset># (N::'a::order multiset)" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2334 | by (fact multiset_order.less_trans) | 
| 58425 | 2335 | |
| 60606 | 2336 | lemma mult_less_not_sym: "M #\<subset># N \<Longrightarrow> \<not> N #\<subset># (M::'a::order multiset)" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2337 | by (fact multiset_order.less_not_sym) | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2338 | |
| 60606 | 2339 | lemma mult_less_asym: "M #\<subset># N \<Longrightarrow> (\<not> P \<Longrightarrow> N #\<subset># (M::'a::order multiset)) \<Longrightarrow> P" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 2340 | 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 | 2341 | |
| 60500 | 2342 | declaration \<open> | 
| 60606 | 2343 | let | 
| 2344 | fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') = | |
| 2345 | let | |
| 2346 | val (maybe_opt, ps) = | |
| 2347 | Nitpick_Model.dest_plain_fun t' | |
| 2348 | ||> op ~~ | |
| 2349 | ||> map (apsnd (snd o HOLogic.dest_number)) | |
| 2350 | fun elems_for t = | |
| 2351 | (case AList.lookup (op =) ps t of | |
| 2352 | SOME n => replicate n t | |
| 2353 | | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]) | |
| 2354 | in | |
| 2355 | (case maps elems_for (all_values elem_T) @ | |
| 61333 | 2356 | (if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of | 
| 60606 | 2357 |               [] => Const (@{const_name zero_class.zero}, T)
 | 
| 2358 | | ts => | |
| 2359 | foldl1 (fn (t1, t2) => | |
| 2360 |                     Const (@{const_name plus_class.plus}, T --> T --> T) $ t1 $ t2)
 | |
| 2361 |                   (map (curry (op $) (Const (@{const_name single}, elem_T --> T))) ts))
 | |
| 2362 | end | |
| 2363 | | multiset_postproc _ _ _ _ t = t | |
| 2364 |   in Nitpick_Model.register_term_postprocessor @{typ "'a multiset"} multiset_postproc end
 | |
| 60500 | 2365 | \<close> | 
| 2366 | ||
| 2367 | ||
| 2368 | subsection \<open>Naive implementation using lists\<close> | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2369 | |
| 60515 | 2370 | code_datatype mset | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2371 | |
| 60606 | 2372 | lemma [code]: "{#} = mset []"
 | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2373 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2374 | |
| 60606 | 2375 | lemma [code]: "{#x#} = mset [x]"
 | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2376 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2377 | |
| 60606 | 2378 | lemma union_code [code]: "mset xs + mset ys = mset (xs @ ys)" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2379 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2380 | |
| 60606 | 2381 | lemma [code]: "image_mset f (mset xs) = mset (map f xs)" | 
| 60515 | 2382 | by (simp add: mset_map) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2383 | |
| 60606 | 2384 | lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)" | 
| 60515 | 2385 | by (simp add: mset_filter) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2386 | |
| 60606 | 2387 | lemma [code]: "mset xs - mset ys = mset (fold remove1 ys xs)" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2388 | 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 | 2389 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2390 | lemma [code]: | 
| 60515 | 2391 | "mset xs #\<inter> mset ys = | 
| 2392 | mset (snd (fold (\<lambda>x (ys, zs). | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2393 | 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 | 2394 | proof - | 
| 60515 | 2395 | have "\<And>zs. mset (snd (fold (\<lambda>x (ys, zs). | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2396 | if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) = | 
| 60515 | 2397 | (mset xs #\<inter> mset ys) + mset zs" | 
| 51623 | 2398 | by (induct xs arbitrary: ys) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2399 | (auto simp add: inter_add_right1 inter_add_right2 ac_simps) | 
| 51623 | 2400 | then show ?thesis by simp | 
| 2401 | qed | |
| 2402 | ||
| 2403 | lemma [code]: | |
| 60515 | 2404 | "mset xs #\<union> mset ys = | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61378diff
changeset | 2405 | mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" | 
| 51623 | 2406 | proof - | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61378diff
changeset | 2407 | have "\<And>zs. mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) = | 
| 60515 | 2408 | (mset xs #\<union> mset ys) + mset zs" | 
| 51623 | 2409 | 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 | 2410 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2411 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2412 | |
| 59813 | 2413 | declare in_multiset_in_set [code_unfold] | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2414 | |
| 60606 | 2415 | lemma [code]: "count (mset xs) x = fold (\<lambda>y. if x = y then Suc else id) xs 0" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2416 | proof - | 
| 60515 | 2417 | have "\<And>n. fold (\<lambda>y. if x = y then Suc else id) xs n = count (mset xs) x + n" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2418 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2419 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2420 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2421 | |
| 60515 | 2422 | declare set_mset_mset [code] | 
| 2423 | ||
| 2424 | declare sorted_list_of_multiset_mset [code] | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2425 | |
| 61585 | 2426 | lemma [code]: \<comment> \<open>not very efficient, but representation-ignorant!\<close> | 
| 60515 | 2427 | "mset_set A = mset (sorted_list_of_set A)" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2428 | apply (cases "finite A") | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2429 | apply simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2430 | apply (induct A rule: finite_induct) | 
| 59813 | 2431 | apply (simp_all add: add.commute) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2432 | done | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2433 | |
| 60515 | 2434 | declare size_mset [code] | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2435 | |
| 58425 | 2436 | 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 | 2437 | "ms_lesseq_impl [] ys = Some (ys \<noteq> [])" | 
| 58425 | 2438 | | "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 | 2439 | None \<Rightarrow> None | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2440 | | 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 | 2441 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2442 | lemma ms_lesseq_impl: "(ms_lesseq_impl xs ys = None \<longleftrightarrow> \<not> mset xs \<subseteq># mset ys) \<and> | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2443 | (ms_lesseq_impl xs ys = Some True \<longleftrightarrow> mset xs \<subset># mset ys) \<and> | 
| 60515 | 2444 | (ms_lesseq_impl xs ys = Some False \<longrightarrow> mset xs = mset ys)" | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2445 | proof (induct xs arbitrary: ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2446 | case (Nil ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2447 | 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 | 2448 | next | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2449 | case (Cons x xs ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2450 | show ?case | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2451 | 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 | 2452 | case None | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2453 | 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 | 2454 |     {
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2455 | assume "mset (x # xs) \<subseteq># mset ys" | 
| 60495 | 2456 | from set_mset_mono[OF this] x have False by simp | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2457 | } note nle = this | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2458 | moreover | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2459 |     {
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2460 | assume "mset (x # xs) \<subset># mset ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2461 | hence "mset (x # xs) \<subseteq># mset ys" by auto | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2462 | 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 | 2463 | } | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2464 | 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 | 2465 | next | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2466 | case (Some res) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2467 | 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 | 2468 | note Some = Some[unfolded res] | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2469 | from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp | 
| 60515 | 2470 |     hence id: "mset ys = mset (ys1 @ ys2) + {#x#}"
 | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2471 | by (auto simp: ac_simps) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2472 | 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 | 2473 | unfolding Some option.simps split | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2474 | unfolding id | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2475 | using Cons[of "ys1 @ ys2"] | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 2476 | unfolding subset_mset_def subseteq_mset_def by auto | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2477 | qed | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2478 | qed | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2479 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2480 | lemma [code]: "mset xs \<subseteq># mset ys \<longleftrightarrow> ms_lesseq_impl xs ys \<noteq> None" | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2481 | 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 | 2482 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2483 | lemma [code]: "mset xs \<subset># mset ys \<longleftrightarrow> ms_lesseq_impl xs ys = Some True" | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2484 | 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 | 2485 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2486 | instantiation multiset :: (equal) equal | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2487 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2488 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2489 | definition | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2490 | [code del]: "HOL.equal A (B :: 'a multiset) \<longleftrightarrow> A = B" | 
| 60515 | 2491 | lemma [code]: "HOL.equal (mset xs) (mset ys) \<longleftrightarrow> ms_lesseq_impl xs ys = Some False" | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2492 | unfolding equal_multiset_def | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 2493 | 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 | 2494 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2495 | instance | 
| 60678 | 2496 | by standard (simp add: equal_multiset_def) | 
| 2497 | ||
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
37107diff
changeset | 2498 | end | 
| 49388 | 2499 | |
| 60606 | 2500 | lemma [code]: "msetsum (mset xs) = listsum xs" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2501 | 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 | 2502 | |
| 60606 | 2503 | lemma [code]: "msetprod (mset xs) = fold times xs 1" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2504 | proof - | 
| 60515 | 2505 | have "\<And>x. fold times xs x = msetprod (mset xs) * x" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2506 | 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 | 2507 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2508 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2509 | |
| 60500 | 2510 | text \<open> | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2511 |   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 | 2512 |   and @{const less_multiset} (multiset order).
 | 
| 60500 | 2513 | \<close> | 
| 2514 | ||
| 2515 | text \<open>Quickcheck generators\<close> | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2516 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2517 | definition (in term_syntax) | 
| 61076 | 2518 | msetify :: "'a::typerep list \<times> (unit \<Rightarrow> Code_Evaluation.term) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2519 | \<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where | 
| 60515 | 2520 |   [code_unfold]: "msetify xs = Code_Evaluation.valtermify mset {\<cdot>} xs"
 | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2521 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2522 | notation fcomp (infixl "\<circ>>" 60) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2523 | notation scomp (infixl "\<circ>\<rightarrow>" 60) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2524 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2525 | instantiation multiset :: (random) random | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2526 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2527 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2528 | definition | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2529 | "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 | 2530 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2531 | instance .. | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2532 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2533 | end | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2534 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2535 | no_notation fcomp (infixl "\<circ>>" 60) | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2536 | 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 | 2537 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2538 | instantiation multiset :: (full_exhaustive) full_exhaustive | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2539 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2540 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2541 | 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 | 2542 | where | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2543 | "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 | 2544 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2545 | instance .. | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2546 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2547 | end | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2548 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2549 | hide_const (open) msetify | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2550 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2551 | |
| 60500 | 2552 | subsection \<open>BNF setup\<close> | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2553 | |
| 57966 | 2554 | definition rel_mset where | 
| 60515 | 2555 | "rel_mset R X Y \<longleftrightarrow> (\<exists>xs ys. mset xs = X \<and> mset ys = Y \<and> list_all2 R xs ys)" | 
| 2556 | ||
| 2557 | lemma mset_zip_take_Cons_drop_twice: | |
| 57966 | 2558 | assumes "length xs = length ys" "j \<le> length xs" | 
| 60515 | 2559 | shows "mset (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) = | 
| 2560 |     mset (zip xs ys) + {#(x, y)#}"
 | |
| 60606 | 2561 | using assms | 
| 57966 | 2562 | proof (induct xs ys arbitrary: x y j rule: list_induct2) | 
| 2563 | case Nil | |
| 2564 | thus ?case | |
| 2565 | by simp | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2566 | next | 
| 57966 | 2567 | case (Cons x xs y ys) | 
| 2568 | thus ?case | |
| 2569 | proof (cases "j = 0") | |
| 2570 | case True | |
| 2571 | thus ?thesis | |
| 2572 | by simp | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2573 | next | 
| 57966 | 2574 | case False | 
| 2575 | then obtain k where k: "j = Suc k" | |
| 60678 | 2576 | by (cases j) simp | 
| 57966 | 2577 | hence "k \<le> length xs" | 
| 2578 | using Cons.prems by auto | |
| 60515 | 2579 | hence "mset (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) = | 
| 2580 |       mset (zip xs ys) + {#(x, y)#}"
 | |
| 57966 | 2581 | by (rule Cons.hyps(2)) | 
| 2582 | thus ?thesis | |
| 2583 | unfolding k by (auto simp: add.commute union_lcomm) | |
| 58425 | 2584 | qed | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2585 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2586 | |
| 60515 | 2587 | lemma ex_mset_zip_left: | 
| 2588 | assumes "length xs = length ys" "mset xs' = mset xs" | |
| 2589 | shows "\<exists>ys'. length ys' = length xs' \<and> mset (zip xs' ys') = mset (zip xs ys)" | |
| 58425 | 2590 | using assms | 
| 57966 | 2591 | proof (induct xs ys arbitrary: xs' rule: list_induct2) | 
| 2592 | case Nil | |
| 2593 | thus ?case | |
| 2594 | by auto | |
| 2595 | next | |
| 2596 | case (Cons x xs y ys xs') | |
| 2597 | obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" | |
| 60515 | 2598 | by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) | 
| 58425 | 2599 | |
| 2600 | def xsa \<equiv> "take j xs' @ drop (Suc j) xs'" | |
| 60515 | 2601 |   have "mset xs' = {#x#} + mset xsa"
 | 
| 57966 | 2602 | 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 | 2603 | by (metis (no_types) ab_semigroup_add_class.add_ac(1) append_take_drop_id Cons_nth_drop_Suc | 
| 60515 | 2604 | mset.simps(2) union_code add.commute) | 
| 2605 | hence ms_x: "mset xsa = mset xs" | |
| 2606 | by (metis Cons.prems add.commute add_right_imp_eq mset.simps(2)) | |
| 57966 | 2607 | then obtain ysa where | 
| 60515 | 2608 | len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" | 
| 57966 | 2609 | using Cons.hyps(2) by blast | 
| 2610 | ||
| 2611 | def ys' \<equiv> "take j ysa @ y # drop j ysa" | |
| 2612 | have xs': "xs' = take j xsa @ x # drop j xsa" | |
| 2613 | 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 | 2614 | by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons | 
| 60515 | 2615 | length_drop size_mset) | 
| 57966 | 2616 | have j_len': "j \<le> length xsa" | 
| 2617 | using j_len xs' xsa_def | |
| 2618 | by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) | |
| 2619 | have "length ys' = length xs'" | |
| 2620 | unfolding ys'_def using Cons.prems len_a ms_x | |
| 60515 | 2621 | by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) | 
| 2622 | moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" | |
| 57966 | 2623 | unfolding xs' ys'_def | 
| 60515 | 2624 | by (rule trans[OF mset_zip_take_Cons_drop_twice]) | 
| 57966 | 2625 | (auto simp: len_a ms_a j_len' add.commute) | 
| 2626 | ultimately show ?case | |
| 2627 | by blast | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2628 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2629 | |
| 57966 | 2630 | lemma list_all2_reorder_left_invariance: | 
| 60515 | 2631 | assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs" | 
| 2632 | shows "\<exists>ys'. list_all2 R xs' ys' \<and> mset ys' = mset ys" | |
| 57966 | 2633 | proof - | 
| 2634 | have len: "length xs = length ys" | |
| 2635 | using rel list_all2_conv_all_nth by auto | |
| 2636 | obtain ys' where | |
| 60515 | 2637 | len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" | 
| 2638 | using len ms_x by (metis ex_mset_zip_left) | |
| 57966 | 2639 | have "list_all2 R xs' ys'" | 
| 60515 | 2640 | using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) | 
| 2641 | moreover have "mset ys' = mset ys" | |
| 2642 | using len len' ms_xy map_snd_zip mset_map by metis | |
| 57966 | 2643 | ultimately show ?thesis | 
| 2644 | by blast | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2645 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2646 | |
| 60515 | 2647 | lemma ex_mset: "\<exists>xs. mset xs = X" | 
| 2648 | by (induct X) (simp, metis mset.simps(2)) | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2649 | |
| 62324 | 2650 | inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool"
 | 
| 2651 | where | |
| 2652 |   "pred_mset P {#}"
 | |
| 2653 | | "\<lbrakk>P a; pred_mset P M\<rbrakk> \<Longrightarrow> pred_mset P (M + {#a#})"
 | |
| 2654 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2655 | bnf "'a multiset" | 
| 57966 | 2656 | map: image_mset | 
| 60495 | 2657 | sets: set_mset | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2658 | bd: natLeq | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2659 |   wits: "{#}"
 | 
| 57966 | 2660 | rel: rel_mset | 
| 62324 | 2661 | pred: pred_mset | 
| 57966 | 2662 | proof - | 
| 2663 | show "image_mset id = id" | |
| 2664 | by (rule image_mset.id) | |
| 60606 | 2665 | show "image_mset (g \<circ> f) = image_mset g \<circ> image_mset f" for f g | 
| 59813 | 2666 | unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality) | 
| 60606 | 2667 | show "(\<And>z. z \<in> set_mset X \<Longrightarrow> f z = g z) \<Longrightarrow> image_mset f X = image_mset g X" for f g X | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2668 | by (induct X) simp_all | 
| 60606 | 2669 | show "set_mset \<circ> image_mset f = op ` f \<circ> set_mset" for f | 
| 57966 | 2670 | by auto | 
| 2671 | show "card_order natLeq" | |
| 2672 | by (rule natLeq_card_order) | |
| 2673 | show "BNF_Cardinal_Arithmetic.cinfinite natLeq" | |
| 2674 | by (rule natLeq_cinfinite) | |
| 60606 | 2675 | show "ordLeq3 (card_of (set_mset X)) natLeq" for X | 
| 57966 | 2676 | by transfer | 
| 2677 | (auto intro!: ordLess_imp_ordLeq simp: finite_iff_ordLess_natLeq[symmetric] multiset_def) | |
| 60606 | 2678 | show "rel_mset R OO rel_mset S \<le> rel_mset (R OO S)" for R S | 
| 57966 | 2679 | unfolding rel_mset_def[abs_def] OO_def | 
| 2680 | apply clarify | |
| 60678 | 2681 | subgoal for X Z Y xs ys' ys zs | 
| 2682 | apply (drule list_all2_reorder_left_invariance [where xs = ys' and ys = zs and xs' = ys]) | |
| 2683 | apply (auto intro: list_all2_trans) | |
| 2684 | done | |
| 60606 | 2685 | done | 
| 2686 | show "rel_mset R = | |
| 62324 | 2687 |     (\<lambda>x y. \<exists>z. set_mset z \<subseteq> {(x, y). R x y} \<and>
 | 
| 2688 | image_mset fst z = x \<and> image_mset snd z = y)" for R | |
| 2689 | unfolding rel_mset_def[abs_def] | |
| 57966 | 2690 | apply (rule ext)+ | 
| 62324 | 2691 | apply safe | 
| 2692 | apply (rule_tac x = "mset (zip xs ys)" in exI; | |
| 2693 | auto simp: in_set_zip list_all2_iff mset_map[symmetric]) | |
| 57966 | 2694 | apply (rename_tac XY) | 
| 60515 | 2695 | apply (cut_tac X = XY in ex_mset) | 
| 57966 | 2696 | apply (erule exE) | 
| 2697 | apply (rename_tac xys) | |
| 2698 | apply (rule_tac x = "map fst xys" in exI) | |
| 60515 | 2699 | apply (auto simp: mset_map) | 
| 57966 | 2700 | apply (rule_tac x = "map snd xys" in exI) | 
| 60515 | 2701 | apply (auto simp: mset_map list_all2I subset_eq zip_map_fst_snd) | 
| 59997 | 2702 | done | 
| 60606 | 2703 |   show "z \<in> set_mset {#} \<Longrightarrow> False" for z
 | 
| 57966 | 2704 | by auto | 
| 62324 | 2705 | show "pred_mset P = (\<lambda>x. Ball (set_mset x) P)" for P | 
| 2706 | proof (intro ext iffI) | |
| 2707 | fix x | |
| 2708 | assume "pred_mset P x" | |
| 2709 | then show "Ball (set_mset x) P" by (induct pred: pred_mset; simp) | |
| 2710 | next | |
| 2711 | fix x | |
| 2712 | assume "Ball (set_mset x) P" | |
| 2713 | then show "pred_mset P x" by (induct x; auto intro: pred_mset.intros) | |
| 2714 | qed | |
| 57966 | 2715 | qed | 
| 2716 | ||
| 60606 | 2717 | inductive rel_mset' | 
| 2718 | where | |
| 57966 | 2719 |   Zero[intro]: "rel_mset' R {#} {#}"
 | 
| 2720 | | Plus[intro]: "\<lbrakk>R a b; rel_mset' R M N\<rbrakk> \<Longrightarrow> rel_mset' R (M + {#a#}) (N + {#b#})"
 | |
| 2721 | ||
| 2722 | lemma rel_mset_Zero: "rel_mset R {#} {#}"
 | |
| 2723 | 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 | 2724 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2725 | declare multiset.count[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2726 | declare Abs_multiset_inverse[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2727 | declare multiset.count_inverse[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2728 | declare union_preserves_multiset[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2729 | |
| 57966 | 2730 | lemma rel_mset_Plus: | 
| 60606 | 2731 | assumes ab: "R a b" | 
| 2732 | and MN: "rel_mset R M N" | |
| 2733 |   shows "rel_mset R (M + {#a#}) (N + {#b#})"
 | |
| 2734 | proof - | |
| 2735 |   have "\<exists>ya. image_mset fst y + {#a#} = image_mset fst ya \<and>
 | |
| 2736 |     image_mset snd y + {#b#} = image_mset snd ya \<and>
 | |
| 2737 |     set_mset ya \<subseteq> {(x, y). R x y}"
 | |
| 2738 |     if "R a b" and "set_mset y \<subseteq> {(x, y). R x y}" for y
 | |
| 2739 |     using that by (intro exI[of _ "y + {#(a,b)#}"]) auto
 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2740 | thus ?thesis | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2741 | using assms | 
| 57966 | 2742 | 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 | 2743 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2744 | |
| 60606 | 2745 | lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \<Longrightarrow> rel_mset R M N" | 
| 60678 | 2746 | by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus) | 
| 57966 | 2747 | |
| 60606 | 2748 | lemma rel_mset_size: "rel_mset R M N \<Longrightarrow> size M = size N" | 
| 60678 | 2749 | 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 | 2750 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2751 | lemma multiset_induct2[case_names empty addL addR]: | 
| 60678 | 2752 |   assumes empty: "P {#} {#}"
 | 
| 2753 |     and addL: "\<And>M N a. P M N \<Longrightarrow> P (M + {#a#}) N"
 | |
| 2754 |     and addR: "\<And>M N a. P M N \<Longrightarrow> P M (N + {#a#})"
 | |
| 2755 | shows "P M N" | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2756 | apply(induct N rule: multiset_induct) | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2757 | 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 | 2758 | 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 | 2759 | done | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2760 | |
| 59949 | 2761 | lemma multiset_induct2_size[consumes 1, case_names empty add]: | 
| 60606 | 2762 | assumes c: "size M = size N" | 
| 2763 |     and empty: "P {#} {#}"
 | |
| 2764 |     and add: "\<And>M N a b. P M N \<Longrightarrow> P (M + {#a#}) (N + {#b#})"
 | |
| 2765 | shows "P M N" | |
| 60678 | 2766 | using c | 
| 2767 | proof (induct M arbitrary: N rule: measure_induct_rule[of size]) | |
| 60606 | 2768 | case (less M) | 
| 2769 | show ?case | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2770 |   proof(cases "M = {#}")
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2771 |     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 | 2772 | thus ?thesis using True empty by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2773 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2774 |     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 | 2775 |     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 | 2776 |     then obtain N1 b where N: "N = N1 + {#b#}" by (metis multi_nonempty_split)
 | 
| 59949 | 2777 | have "size M1 = size N1" using less.prems unfolding M N by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2778 | 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 | 2779 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2780 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2781 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2782 | lemma msed_map_invL: | 
| 60606 | 2783 |   assumes "image_mset f (M + {#a#}) = N"
 | 
| 2784 |   shows "\<exists>N1. N = N1 + {#f a#} \<and> image_mset f M = N1"
 | |
| 2785 | proof - | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2786 | have "f a \<in># N" | 
| 60606 | 2787 |     using assms multiset.set_map[of f "M + {#a#}"] by auto
 | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2788 |   then obtain N1 where N: "N = N1 + {#f a#}" using multi_member_split by metis
 | 
| 57966 | 2789 | 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 | 2790 | thus ?thesis using N by blast | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2791 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2792 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2793 | lemma msed_map_invR: | 
| 60606 | 2794 |   assumes "image_mset f M = N + {#b#}"
 | 
| 2795 |   shows "\<exists>M1 a. M = M1 + {#a#} \<and> f a = b \<and> image_mset f M1 = N"
 | |
| 2796 | proof - | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2797 | obtain a where a: "a \<in># M" and fa: "f a = b" | 
| 60606 | 2798 | using multiset.set_map[of f M] unfolding assms | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2799 | by (metis image_iff union_single_eq_member) | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2800 |   then obtain M1 where M: "M = M1 + {#a#}" using multi_member_split by metis
 | 
| 57966 | 2801 | 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 | 2802 | thus ?thesis using M fa by blast | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2803 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2804 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2805 | lemma msed_rel_invL: | 
| 60606 | 2806 |   assumes "rel_mset R (M + {#a#}) N"
 | 
| 2807 |   shows "\<exists>N1 b. N = N1 + {#b#} \<and> R a b \<and> rel_mset R M N1"
 | |
| 2808 | proof - | |
| 57966 | 2809 |   obtain K where KM: "image_mset fst K = M + {#a#}"
 | 
| 60606 | 2810 |     and KN: "image_mset snd K = N" and sK: "set_mset K \<subseteq> {(a, b). R a b}"
 | 
| 2811 | using assms | |
| 2812 | 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 | 2813 |   obtain K1 ab where K: "K = K1 + {#ab#}" and a: "fst ab = a"
 | 
| 60606 | 2814 | and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto | 
| 57966 | 2815 |   obtain N1 where N: "N = N1 + {#snd ab#}" and K1N1: "image_mset snd K1 = N1"
 | 
| 60606 | 2816 | using msed_map_invL[OF KN[unfolded K]] by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2817 | have Rab: "R a (snd ab)" using sK a unfolding K by auto | 
| 57966 | 2818 | have "rel_mset R M N1" using sK K1M K1N1 | 
| 60606 | 2819 | 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 | 2820 | thus ?thesis using N Rab by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2821 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2822 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2823 | lemma msed_rel_invR: | 
| 60606 | 2824 |   assumes "rel_mset R M (N + {#b#})"
 | 
| 2825 |   shows "\<exists>M1 a. M = M1 + {#a#} \<and> R a b \<and> rel_mset R M1 N"
 | |
| 2826 | proof - | |
| 57966 | 2827 |   obtain K where KN: "image_mset snd K = N + {#b#}"
 | 
| 60606 | 2828 |     and KM: "image_mset fst K = M" and sK: "set_mset K \<subseteq> {(a, b). R a b}"
 | 
| 2829 | using assms | |
| 2830 | 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 | 2831 |   obtain K1 ab where K: "K = K1 + {#ab#}" and b: "snd ab = b"
 | 
| 60606 | 2832 | and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto | 
| 57966 | 2833 |   obtain M1 where M: "M = M1 + {#fst ab#}" and K1M1: "image_mset fst K1 = M1"
 | 
| 60606 | 2834 | using msed_map_invL[OF KM[unfolded K]] by auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2835 | have Rab: "R (fst ab) b" using sK b unfolding K by auto | 
| 57966 | 2836 | have "rel_mset R M1 N" using sK K1N K1M1 | 
| 60606 | 2837 | 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 | 2838 | thus ?thesis using M Rab by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2839 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2840 | |
| 57966 | 2841 | lemma rel_mset_imp_rel_mset': | 
| 60606 | 2842 | assumes "rel_mset R M N" | 
| 2843 | shows "rel_mset' R M N" | |
| 59949 | 2844 | using assms proof(induct M arbitrary: N rule: measure_induct_rule[of size]) | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2845 | case (less M) | 
| 59949 | 2846 | have c: "size M = size N" using rel_mset_size[OF less.prems] . | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2847 | show ?case | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2848 |   proof(cases "M = {#}")
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2849 |     case True hence "N = {#}" using c by simp
 | 
| 57966 | 2850 | 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 | 2851 | next | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2852 |     case False then obtain M1 a where M: "M = M1 + {#a#}" by (metis multi_nonempty_split)
 | 
| 57966 | 2853 |     obtain N1 b where N: "N = N1 + {#b#}" and R: "R a b" and ms: "rel_mset R M1 N1"
 | 
| 60606 | 2854 | using msed_rel_invL[OF less.prems[unfolded M]] by auto | 
| 57966 | 2855 | have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp | 
| 2856 | 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 | 2857 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2858 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2859 | |
| 60606 | 2860 | lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N" | 
| 60678 | 2861 | using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto | 
| 57966 | 2862 | |
| 60613 | 2863 | text \<open>The main end product for @{const rel_mset}: inductive characterization:\<close>
 | 
| 61337 | 2864 | lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] = | 
| 60606 | 2865 | 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 | 2866 | |
| 56656 | 2867 | |
| 60500 | 2868 | subsection \<open>Size setup\<close> | 
| 56656 | 2869 | |
| 57966 | 2870 | lemma multiset_size_o_map: "size_multiset g \<circ> image_mset f = size_multiset (g \<circ> f)" | 
| 60678 | 2871 | apply (rule ext) | 
| 2872 | subgoal for x by (induct x) auto | |
| 2873 | done | |
| 56656 | 2874 | |
| 60500 | 2875 | setup \<open> | 
| 60606 | 2876 |   BNF_LFP_Size.register_size_global @{type_name multiset} @{const_name size_multiset}
 | 
| 62082 | 2877 |     @{thm size_multiset_overloaded_def}
 | 
| 60606 | 2878 |     @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single
 | 
| 2879 | size_union} | |
| 2880 |     @{thms multiset_size_o_map}
 | |
| 60500 | 2881 | \<close> | 
| 56656 | 2882 | |
| 2883 | hide_const (open) wcount | |
| 2884 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 2885 | end |