| author | wenzelm | 
| Wed, 11 Sep 2024 20:05:09 +0200 | |
| changeset 80858 | a392351d1ed4 | 
| parent 80786 | 70076ba563d2 | 
| child 80914 | d97fdabd9e2b | 
| 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 | |
| 74803 | 7 | Author: Martin Desharnais, MPI-INF Saarbruecken | 
| 10249 | 8 | *) | 
| 9 | ||
| 65048 | 10 | section \<open>(Finite) Multisets\<close> | 
| 10249 | 11 | |
| 15131 | 12 | theory Multiset | 
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 13 | imports Cancellation | 
| 15131 | 14 | begin | 
| 10249 | 15 | |
| 60500 | 16 | subsection \<open>The type of multisets\<close> | 
| 10249 | 17 | |
| 73270 | 18 | typedef 'a multiset = \<open>{f :: 'a \<Rightarrow> nat. finite {x. f x > 0}}\<close>
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 19 | morphisms count Abs_multiset | 
| 10249 | 20 | proof | 
| 73270 | 21 |   show \<open>(\<lambda>x. 0::nat) \<in> {f. finite {x. f x > 0}}\<close>
 | 
| 22 | by simp | |
| 10249 | 23 | qed | 
| 24 | ||
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 25 | setup_lifting type_definition_multiset | 
| 19086 | 26 | |
| 73270 | 27 | lemma count_Abs_multiset: | 
| 28 |   \<open>count (Abs_multiset f) = f\<close> if \<open>finite {x. f x > 0}\<close>
 | |
| 29 | by (rule Abs_multiset_inverse) (simp add: that) | |
| 30 | ||
| 60606 | 31 | lemma multiset_eq_iff: "M = N \<longleftrightarrow> (\<forall>a. count M a = count N a)" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 32 | by (simp only: count_inject [symmetric] fun_eq_iff) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 33 | |
| 60606 | 34 | lemma multiset_eqI: "(\<And>x. count A x = count B x) \<Longrightarrow> A = B" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 35 | using multiset_eq_iff by auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 36 | |
| 69593 | 37 | text \<open>Preservation of the representing set \<^term>\<open>multiset\<close>.\<close> | 
| 60606 | 38 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 39 | lemma diff_preserves_multiset: | 
| 73270 | 40 |   \<open>finite {x. 0 < M x - N x}\<close> if \<open>finite {x. 0 < M x}\<close> for M N :: \<open>'a \<Rightarrow> nat\<close>
 | 
| 41 | using that by (rule rev_finite_subset) auto | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 42 | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 43 | lemma filter_preserves_multiset: | 
| 73270 | 44 |   \<open>finite {x. 0 < (if P x then M x else 0)}\<close> if \<open>finite {x. 0 < M x}\<close> for M N :: \<open>'a \<Rightarrow> nat\<close>
 | 
| 45 | using that by (rule rev_finite_subset) auto | |
| 46 | ||
| 47 | lemmas in_multiset = diff_preserves_multiset filter_preserves_multiset | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 48 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 49 | |
| 60500 | 50 | subsection \<open>Representing multisets\<close> | 
| 51 | ||
| 52 | text \<open>Multiset enumeration\<close> | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 53 | |
| 48008 | 54 | instantiation multiset :: (type) cancel_comm_monoid_add | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 55 | begin | 
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 56 | |
| 73393 | 57 | lift_definition zero_multiset :: \<open>'a multiset\<close> | 
| 58 | is \<open>\<lambda>a. 0\<close> | |
| 73270 | 59 | by simp | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 60 | |
| 73393 | 61 | abbreviation empty_mset :: \<open>'a multiset\<close> (\<open>{#}\<close>)
 | 
| 62 | where \<open>empty_mset \<equiv> 0\<close> | |
| 63 | ||
| 64 | lift_definition plus_multiset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> | |
| 65 | is \<open>\<lambda>M N a. M a + N a\<close> | |
| 73270 | 66 | by simp | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 67 | |
| 73393 | 68 | lift_definition minus_multiset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> | 
| 69 | is \<open>\<lambda>M N a. M a - N a\<close> | |
| 73270 | 70 | by (rule diff_preserves_multiset) | 
| 59815 
cce82e360c2f
explicit commutative additive inverse operation;
 haftmann parents: 
59813diff
changeset | 71 | |
| 48008 | 72 | instance | 
| 73270 | 73 | by (standard; transfer) (simp_all add: fun_eq_iff) | 
| 25571 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 74 | |
| 
c9e39eafc7a0
instantiation target rather than legacy instance
 haftmann parents: 
25507diff
changeset | 75 | end | 
| 10249 | 76 | |
| 63195 | 77 | context | 
| 78 | begin | |
| 79 | ||
| 80 | qualified definition is_empty :: "'a multiset \<Rightarrow> bool" where | |
| 81 |   [code_abbrev]: "is_empty A \<longleftrightarrow> A = {#}"
 | |
| 82 | ||
| 83 | end | |
| 84 | ||
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 85 | lemma add_mset_in_multiset: | 
| 73270 | 86 |   \<open>finite {x. 0 < (if x = a then Suc (M x) else M x)}\<close>
 | 
| 87 |   if \<open>finite {x. 0 < M x}\<close>
 | |
| 88 | using that by (simp add: flip: insert_Collect) | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 89 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 90 | lift_definition add_mset :: "'a \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" is | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 91 | "\<lambda>a M b. if b = a then Suc (M b) else M b" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 92 | by (rule add_mset_in_multiset) | 
| 15869 | 93 | |
| 80786 
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
 wenzelm parents: 
80768diff
changeset | 94 | nonterminal multiset_args | 
| 26145 | 95 | syntax | 
| 80786 
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
 wenzelm parents: 
80768diff
changeset | 96 |   "" :: "'a \<Rightarrow> multiset_args"  ("_")
 | 
| 
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
 wenzelm parents: 
80768diff
changeset | 97 |   "_multiset_args" :: "'a \<Rightarrow> multiset_args \<Rightarrow> multiset_args"  ("_,/ _")
 | 
| 
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
 wenzelm parents: 
80768diff
changeset | 98 |   "_multiset" :: "multiset_args \<Rightarrow> 'a multiset"    ("{#(_)#}")
 | 
| 80768 | 99 | syntax_consts | 
| 80786 
70076ba563d2
more specific "args" syntax, to support more markup for syntax consts;
 wenzelm parents: 
80768diff
changeset | 100 | "_multiset_args" "_multiset" == add_mset | 
| 25507 | 101 | translations | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 102 |   "{#x, xs#}" == "CONST add_mset x {#xs#}"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 103 |   "{#x#}" == "CONST add_mset x {#}"
 | 
| 25507 | 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_empty [simp]: "count {#} a = 0"
 | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 106 | by (simp add: zero_multiset.rep_eq) | 
| 10249 | 107 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 108 | lemma count_add_mset [simp]: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 109 | "count (add_mset b A) a = (if b = a then Suc (count A a) else count A a)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 110 | by (simp add: add_mset.rep_eq) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 111 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 112 | lemma count_single: "count {#b#} a = (if b = a then 1 else 0)"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 113 | by simp | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 114 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 115 | lemma | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 116 |   add_mset_not_empty [simp]: \<open>add_mset a A \<noteq> {#}\<close> and
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 117 |   empty_not_add_mset [simp]: "{#} \<noteq> add_mset a A"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 118 | by (auto simp: multiset_eq_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 119 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 120 | lemma add_mset_add_mset_same_iff [simp]: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 121 | "add_mset a A = add_mset a B \<longleftrightarrow> A = B" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 122 | by (auto simp: multiset_eq_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 123 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 124 | lemma add_mset_commute: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 125 | "add_mset x (add_mset y M) = add_mset y (add_mset x M)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 126 | by (auto simp: multiset_eq_iff) | 
| 29901 | 127 | |
| 10249 | 128 | |
| 60500 | 129 | subsection \<open>Basic operations\<close> | 
| 130 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 131 | subsubsection \<open>Conversion to set and membership\<close> | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 132 | |
| 73393 | 133 | definition set_mset :: \<open>'a multiset \<Rightarrow> 'a set\<close> | 
| 134 |   where \<open>set_mset M = {x. count M x > 0}\<close>
 | |
| 135 | ||
| 136 | abbreviation member_mset :: \<open>'a \<Rightarrow> 'a multiset \<Rightarrow> bool\<close> | |
| 137 | where \<open>member_mset a M \<equiv> a \<in> set_mset M\<close> | |
| 62537 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 138 | |
| 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 139 | notation | 
| 73393 | 140 | member_mset (\<open>'(\<in>#')\<close>) and | 
| 73394 | 141 | member_mset (\<open>(_/ \<in># _)\<close> [50, 51] 50) | 
| 62537 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 142 | |
| 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 143 | notation (ASCII) | 
| 73393 | 144 | member_mset (\<open>'(:#')\<close>) and | 
| 73394 | 145 | member_mset (\<open>(_/ :# _)\<close> [50, 51] 50) | 
| 73393 | 146 | |
| 147 | abbreviation not_member_mset :: \<open>'a \<Rightarrow> 'a multiset \<Rightarrow> bool\<close> | |
| 148 | where \<open>not_member_mset a M \<equiv> a \<notin> set_mset M\<close> | |
| 62537 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 149 | |
| 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 150 | notation | 
| 73393 | 151 | not_member_mset (\<open>'(\<notin>#')\<close>) and | 
| 73394 | 152 | not_member_mset (\<open>(_/ \<notin># _)\<close> [50, 51] 50) | 
| 62537 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 153 | |
| 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 154 | notation (ASCII) | 
| 73393 | 155 | not_member_mset (\<open>'(~:#')\<close>) and | 
| 73394 | 156 | not_member_mset (\<open>(_/ ~:# _)\<close> [50, 51] 50) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 157 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 158 | context | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 159 | begin | 
| 
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 | 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 | 162 | where "Ball M \<equiv> Set.Ball (set_mset M)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 163 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 164 | 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 | 165 | where "Bex M \<equiv> Set.Bex (set_mset M)" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 166 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 167 | end | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 168 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 169 | syntax | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 170 |   "_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 | 171 |   "_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 | 172 | |
| 62537 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 173 | syntax (ASCII) | 
| 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 174 |   "_MBall"       :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<forall>_:#_./ _)" [0, 0, 10] 10)
 | 
| 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 175 |   "_MBex"        :: "pttrn \<Rightarrow> 'a set \<Rightarrow> bool \<Rightarrow> bool"      ("(3\<exists>_:#_./ _)" [0, 0, 10] 10)
 | 
| 
7a9aa69f9b38
syntax for multiset membership modelled after syntax for set membership
 haftmann parents: 
62430diff
changeset | 176 | |
| 80768 | 177 | syntax_consts | 
| 178 | "_MBall" \<rightleftharpoons> Multiset.Ball and | |
| 179 | "_MBex" \<rightleftharpoons> Multiset.Bex | |
| 180 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 181 | translations | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 182 | "\<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 | 183 | "\<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 | 184 | |
| 71917 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 nipkow parents: 
71398diff
changeset | 185 | print_translation \<open> | 
| 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 nipkow parents: 
71398diff
changeset | 186 | [Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\<open>Multiset.Ball\<close> \<^syntax_const>\<open>_MBall\<close>, | 
| 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 nipkow parents: 
71398diff
changeset | 187 | Syntax_Trans.preserve_binder_abs2_tr' \<^const_syntax>\<open>Multiset.Bex\<close> \<^syntax_const>\<open>_MBex\<close>] | 
| 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 nipkow parents: 
71398diff
changeset | 188 | \<close> \<comment> \<open>to avoid eta-contraction of body\<close> | 
| 
4c5778d8a53d
should have been copied across from Set.thy as well for better printing
 nipkow parents: 
71398diff
changeset | 189 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 190 | lemma count_eq_zero_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 191 | "count M x = 0 \<longleftrightarrow> x \<notin># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 192 | by (auto simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 193 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 194 | lemma not_in_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 195 | "x \<notin># M \<longleftrightarrow> count M x = 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 196 | by (auto simp add: count_eq_zero_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 197 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 198 | lemma count_greater_zero_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 199 | "count M x > 0 \<longleftrightarrow> x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 200 | by (auto simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 201 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 202 | lemma count_inI: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 203 | assumes "count M x = 0 \<Longrightarrow> False" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 204 | shows "x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 205 | proof (rule ccontr) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 206 | assume "x \<notin># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 207 | 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 | 208 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 209 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 210 | lemma in_countE: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 211 | assumes "x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 212 | obtains n where "count M x = Suc n" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 213 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 214 | from assms have "count M x > 0" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 215 | then obtain n where "count M x = Suc n" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 216 | using gr0_conv_Suc by blast | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 217 | with that show thesis . | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 218 | qed | 
| 
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_greater_eq_Suc_zero_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 221 | "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 | 222 | by (simp add: Suc_le_eq) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 223 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 224 | lemma count_greater_eq_one_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 225 | "count M x \<ge> 1 \<longleftrightarrow> x \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 226 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 227 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 228 | lemma set_mset_empty [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 229 |   "set_mset {#} = {}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 230 | by (simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 231 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 232 | lemma set_mset_single: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 233 |   "set_mset {#b#} = {b}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 234 | by (simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 235 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 236 | lemma set_mset_eq_empty_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 237 |   "set_mset M = {} \<longleftrightarrow> M = {#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 238 | 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 | 239 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 240 | lemma finite_set_mset [iff]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 241 | "finite (set_mset M)" | 
| 73270 | 242 | using count [of M] by simp | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 243 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 244 | lemma set_mset_add_mset_insert [simp]: \<open>set_mset (add_mset a A) = insert a (set_mset A)\<close> | 
| 68406 | 245 | by (auto simp flip: count_greater_eq_Suc_zero_iff split: if_splits) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 246 | |
| 63924 | 247 | lemma multiset_nonemptyE [elim]: | 
| 248 |   assumes "A \<noteq> {#}"
 | |
| 249 | obtains x where "x \<in># A" | |
| 250 | proof - | |
| 251 | have "\<exists>x. x \<in># A" by (rule ccontr) (insert assms, auto) | |
| 252 | with that show ?thesis by blast | |
| 253 | qed | |
| 254 | ||
| 79800 | 255 | lemma count_gt_imp_in_mset: "count M x > n \<Longrightarrow> x \<in># M" | 
| 256 | using count_greater_zero_iff by fastforce | |
| 257 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 258 | |
| 60500 | 259 | subsubsection \<open>Union\<close> | 
| 10249 | 260 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 261 | lemma count_union [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 262 | "count (M + N) a = count M a + count N a" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 263 | by (simp add: plus_multiset.rep_eq) | 
| 10249 | 264 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 265 | lemma set_mset_union [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 266 | "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 | 267 | 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 | 268 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 269 | lemma union_mset_add_mset_left [simp]: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 270 | "add_mset a A + B = add_mset a (A + B)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 271 | by (auto simp: multiset_eq_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 272 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 273 | lemma union_mset_add_mset_right [simp]: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 274 | "A + add_mset a B = add_mset a (A + B)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 275 | by (auto simp: multiset_eq_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 276 | |
| 80525 | 277 | (* TODO: reverse arguments to prevent unfolding loop *) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 278 | lemma add_mset_add_single: \<open>add_mset a A = A + {#a#}\<close>
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 279 | by (subst union_mset_add_mset_right, subst add.comm_neutral) standard | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 280 | |
| 10249 | 281 | |
| 60500 | 282 | subsubsection \<open>Difference\<close> | 
| 10249 | 283 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 284 | instance multiset :: (type) comm_monoid_diff | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 285 | by standard (transfer; simp add: fun_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 286 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 287 | lemma count_diff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 288 | "count (M - N) a = count M a - count N a" | 
| 47429 
ec64d94cbf9c
multiset operations are defined with lift_definitions;
 bulwahn parents: 
47308diff
changeset | 289 | 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 | 290 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 291 | lemma add_mset_diff_bothsides: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 292 | \<open>add_mset a M - add_mset a A = M - A\<close> | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 293 | by (auto simp: multiset_eq_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 294 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 295 | lemma in_diff_count: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 296 | "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 | 297 | by (simp add: set_mset_def) | 
| 
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 count_in_diffI: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 300 | 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 | 301 | shows "x \<in># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 302 | proof (rule ccontr) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 303 | assume "x \<notin># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 304 | 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 | 305 | by (simp add: in_diff_count not_less) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 306 | with assms show False by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 307 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 308 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 309 | lemma in_diff_countE: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 310 | assumes "x \<in># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 311 | 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 | 312 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 313 | 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 | 314 | 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 | 315 | 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 | 316 | using less_iff_Suc_add by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 317 | with that show thesis . | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 318 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 319 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 320 | lemma in_diffD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 321 | assumes "a \<in># M - N" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 322 | shows "a \<in># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 323 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 324 | have "0 \<le> count N a" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 325 | 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 | 326 | by (simp add: in_diff_count) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 327 | finally show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 328 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 329 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 330 | lemma set_mset_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 331 |   "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 | 332 | by (simp add: set_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 333 | |
| 17161 | 334 | lemma diff_empty [simp]: "M - {#} = M \<and> {#} - M = {#}"
 | 
| 52289 | 335 | by rule (fact Groups.diff_zero, fact Groups.zero_diff) | 
| 36903 | 336 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 337 | lemma diff_cancel: "A - A = {#}"
 | 
| 52289 | 338 | by (fact Groups.diff_cancel) | 
| 10249 | 339 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 340 | lemma diff_union_cancelR: "M + N - N = (M::'a multiset)" | 
| 52289 | 341 | by (fact add_diff_cancel_right') | 
| 10249 | 342 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 343 | lemma diff_union_cancelL: "N + M - N = (M::'a multiset)" | 
| 52289 | 344 | 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 | 345 | |
| 52289 | 346 | lemma diff_right_commute: | 
| 60606 | 347 | fixes M N Q :: "'a multiset" | 
| 348 | shows "M - N - Q = M - Q - N" | |
| 52289 | 349 | by (fact diff_right_commute) | 
| 350 | ||
| 351 | lemma diff_add: | |
| 60606 | 352 | fixes M N Q :: "'a multiset" | 
| 353 | shows "M - (N + Q) = M - N - Q" | |
| 52289 | 354 | by (rule sym) (fact diff_diff_add) | 
| 58425 | 355 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 356 | lemma insert_DiffM [simp]: "x \<in># M \<Longrightarrow> add_mset x (M - {#x#}) = M"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 357 | 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 | 358 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 359 | lemma insert_DiffM2: "x \<in># M \<Longrightarrow> (M - {#x#}) + {#x#} = M"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 360 | by simp | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 361 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 362 | lemma diff_union_swap: "a \<noteq> b \<Longrightarrow> add_mset b (M - {#a#}) = add_mset b M - {#a#}"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 363 | 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 | 364 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 365 | lemma diff_add_mset_swap [simp]: "b \<notin># A \<Longrightarrow> add_mset b M - A = add_mset b (M - A)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 366 | by (auto simp add: multiset_eq_iff simp: not_in_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 367 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 368 | lemma diff_union_swap2 [simp]: "y \<in># M \<Longrightarrow> add_mset x M - {#y#} = add_mset x (M - {#y#})"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 369 | by (metis add_mset_diff_bothsides diff_union_swap diff_zero insert_DiffM) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 370 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 371 | lemma diff_diff_add_mset [simp]: "(M::'a multiset) - N - P = M - (N + P)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 372 | by (rule diff_diff_add) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 373 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 374 | lemma diff_union_single_conv: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 375 |   "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 | 376 | by (simp add: multiset_eq_iff Suc_le_eq) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 377 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 378 | lemma mset_add [elim?]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 379 | assumes "a \<in># A" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 380 | obtains B where "A = add_mset a B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 381 | proof - | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 382 |   from assms have "A = add_mset a (A - {#a#})"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 383 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 384 | with that show thesis . | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 385 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 386 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 387 | lemma union_iff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 388 | "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 | 389 | by auto | 
| 26143 
314c0bcb7df7
Added useful general lemmas from the work with the HeapMonad
 bulwahn parents: 
26033diff
changeset | 390 | |
| 77987 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 391 | lemma count_minus_inter_lt_count_minus_inter_iff: | 
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 392 | "count (M2 - M1) y < count (M1 - M2) y \<longleftrightarrow> y \<in># M1 - M2" | 
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 393 | by (meson count_greater_zero_iff gr_implies_not_zero in_diff_count leI order.strict_trans2 | 
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 394 | order_less_asym) | 
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 395 | |
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 396 | lemma minus_inter_eq_minus_inter_iff: | 
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 397 | "(M1 - M2) = (M2 - M1) \<longleftrightarrow> set_mset (M1 - M2) = set_mset (M2 - M1)" | 
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 398 | by (metis add.commute count_diff count_eq_zero_iff diff_add_zero in_diff_countE multiset_eq_iff) | 
| 
0f7dc48d8b7f
added lemmas count_minus_inter_lt_count_minus_inter_iff and minus_inter_eq_minus_inter_iff
 desharna parents: 
77832diff
changeset | 399 | |
| 10249 | 400 | |
| 66425 | 401 | subsubsection \<open>Min and Max\<close> | 
| 402 | ||
| 403 | abbreviation Min_mset :: "'a::linorder multiset \<Rightarrow> 'a" where | |
| 404 | "Min_mset m \<equiv> Min (set_mset m)" | |
| 405 | ||
| 406 | abbreviation Max_mset :: "'a::linorder multiset \<Rightarrow> 'a" where | |
| 407 | "Max_mset m \<equiv> Max (set_mset m)" | |
| 408 | ||
| 79800 | 409 | lemma | 
| 410 |   Min_in_mset: "M \<noteq> {#} \<Longrightarrow> Min_mset M \<in># M" and
 | |
| 411 |   Max_in_mset: "M \<noteq> {#} \<Longrightarrow> Max_mset M \<in># M"
 | |
| 412 | by simp+ | |
| 413 | ||
| 66425 | 414 | |
| 60500 | 415 | 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 | 416 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 417 | 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 | 418 | 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 | 419 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 420 | 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 | 421 | 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 | 422 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 423 | 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 | 424 | 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 | 425 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 426 | lemma multi_self_add_other_not_self [simp]: "M = add_mset x M \<longleftrightarrow> False" | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 427 | 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 | 428 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 429 | lemma add_mset_remove_trivial [simp]: \<open>add_mset x M - {#x#} = M\<close>
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 430 | by (auto simp: multiset_eq_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 431 | |
| 60606 | 432 | 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 | 433 | 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 | 434 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 435 | lemma diff_single_eq_union: "x \<in># M \<Longrightarrow> M - {#x#} = N \<longleftrightarrow> M = add_mset x N"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 436 | by auto | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 437 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 438 | lemma union_single_eq_diff: "add_mset x M = N \<Longrightarrow> M = N - {#x#}"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 439 | unfolding add_mset_add_single[of _ M] by (fact add_implies_diff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 440 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 441 | lemma union_single_eq_member: "add_mset x M = N \<Longrightarrow> x \<in># N" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 442 | 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 | 443 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 444 | lemma add_mset_remove_trivial_If: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 445 |   "add_mset a (N - {#a#}) = (if a \<in># N then N else add_mset a N)"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 446 | by (simp add: diff_single_trivial) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 447 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 448 | lemma add_mset_remove_trivial_eq: \<open>N = add_mset a (N - {#a#}) \<longleftrightarrow> a \<in># N\<close>
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 449 | by (auto simp: add_mset_remove_trivial_If) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 450 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 451 | lemma union_is_single: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 452 |   "M + N = {#a#} \<longleftrightarrow> M = {#a#} \<and> N = {#} \<or> M = {#} \<and> N = {#a#}"
 | 
| 60606 | 453 | (is "?lhs = ?rhs") | 
| 46730 | 454 | proof | 
| 60606 | 455 | show ?lhs if ?rhs using that by auto | 
| 456 | show ?rhs if ?lhs | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 457 | 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 | 458 | 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 | 459 | |
| 60606 | 460 | 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 | 461 |   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 | 462 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 463 | lemma add_eq_conv_diff: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 464 |   "add_mset a M = add_mset b N \<longleftrightarrow> M = N \<and> a = b \<or> M = add_mset b (N - {#a#}) \<and> N = add_mset a (M - {#b#})"
 | 
| 60606 | 465 | (is "?lhs \<longleftrightarrow> ?rhs") | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44339diff
changeset | 466 | (* 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 | 467 | proof | 
| 60606 | 468 | show ?lhs if ?rhs | 
| 469 | using that | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 470 | by (auto simp add: add_mset_commute[of a b]) | 
| 60606 | 471 | 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 | 472 | proof (cases "a = b") | 
| 60500 | 473 | 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 | 474 | 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 | 475 | case False | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 476 | from \<open>?lhs\<close> have "a \<in># add_mset b N" by (rule union_single_eq_member) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 477 | with False have "a \<in># N" by auto | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 478 |     moreover from \<open>?lhs\<close> have "M = add_mset b N - {#a#}" by (rule union_single_eq_diff)
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 479 | moreover note False | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 480 |     ultimately show ?thesis by (auto simp add: diff_right_commute [of _ "{#a#}"])
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 481 | 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 | 482 | 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 | 483 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 484 | lemma add_mset_eq_single [iff]: "add_mset b M = {#a#} \<longleftrightarrow> b = a \<and> M = {#}"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 485 | by (auto simp: add_eq_conv_diff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 486 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 487 | lemma single_eq_add_mset [iff]: "{#a#} = add_mset b M \<longleftrightarrow> b = a \<and> M = {#}"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 488 | by (auto simp: add_eq_conv_diff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 489 | |
| 58425 | 490 | lemma insert_noteq_member: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 491 | assumes BC: "add_mset b B = add_mset c C" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 492 | 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 | 493 | 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 | 494 | proof - | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 495 | have "c \<in># add_mset c C" by simp | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 496 |   have nc: "\<not> c \<in># {#b#}" using bnotc by simp
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 497 | then have "c \<in># add_mset b B" using BC by simp | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 498 | 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 | 499 | 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 | 500 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 501 | lemma add_eq_conv_ex: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 502 | "(add_mset a M = add_mset b N) = | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 503 | (M = N \<and> a = b \<or> (\<exists>K. M = add_mset b K \<and> N = add_mset a K))" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 504 | 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 | 505 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 506 | lemma multi_member_split: "x \<in># M \<Longrightarrow> \<exists>A. M = add_mset x A" | 
| 60678 | 507 |   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 | 508 | |
| 58425 | 509 | lemma multiset_add_sub_el_shuffle: | 
| 60606 | 510 | assumes "c \<in># B" | 
| 511 | and "b \<noteq> c" | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 512 |   shows "add_mset b (B - {#c#}) = add_mset b B - {#c#}"
 | 
| 58098 | 513 | proof - | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 514 | from \<open>c \<in># B\<close> obtain A where B: "B = add_mset c A" | 
| 58098 | 515 | by (blast dest: multi_member_split) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 516 |   have "add_mset b A = add_mset c (add_mset b A) - {#c#}" by simp
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 517 |   then have "add_mset b A = add_mset b (add_mset c A) - {#c#}"
 | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 518 | by (simp add: \<open>b \<noteq> c\<close>) | 
| 58098 | 519 | then show ?thesis using B by simp | 
| 520 | qed | |
| 521 | ||
| 64418 | 522 | lemma add_mset_eq_singleton_iff[iff]: | 
| 523 |   "add_mset x M = {#y#} \<longleftrightarrow> M = {#} \<and> x = y"
 | |
| 524 | by auto | |
| 525 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 526 | |
| 60500 | 527 | 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 | 528 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 529 | definition subseteq_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subseteq>#" 50) | 
| 65466 | 530 | where "A \<subseteq># B \<longleftrightarrow> (\<forall>a. count A a \<le> count B a)" | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 531 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 532 | definition subset_mset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" (infix "\<subset>#" 50) | 
| 65466 | 533 | where "A \<subset># B \<longleftrightarrow> A \<subseteq># B \<and> A \<noteq> B" | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 534 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 535 | 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 | 536 | where "supseteq_mset A B \<equiv> B \<subseteq># A" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 537 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 538 | 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 | 539 | where "supset_mset A B \<equiv> B \<subset># A" | 
| 62208 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 540 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 541 | notation (input) | 
| 62208 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 542 | subseteq_mset (infix "\<le>#" 50) and | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 543 | 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 | 544 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 545 | notation (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 546 | subseteq_mset (infix "<=#" 50) and | 
| 62208 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 547 | subset_mset (infix "<#" 50) and | 
| 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 548 | supseteq_mset (infix ">=#" 50) and | 
| 
ad43b3ab06e4
added 'supset' variants for new '<#' etc. symbols on multisets
 blanchet parents: 
62082diff
changeset | 549 | supset_mset (infix ">#" 50) | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 550 | |
| 73411 | 551 | global_interpretation subset_mset: ordering \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> | 
| 552 | by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order.trans order.antisym) | |
| 553 | ||
| 73451 | 554 | interpretation subset_mset: ordered_ab_semigroup_add_imp_le \<open>(+)\<close> \<open>(-)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> | 
| 60678 | 555 | by standard (auto simp add: subset_mset_def subseteq_mset_def multiset_eq_iff intro: order_trans antisym) | 
| 64585 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 haftmann parents: 
64531diff
changeset | 556 | \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 557 | |
| 67398 | 558 | interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "(+)" 0 "(-)" "(\<subseteq>#)" "(\<subset>#)" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 559 | by standard | 
| 64585 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 haftmann parents: 
64531diff
changeset | 560 | \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 561 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 562 | lemma mset_subset_eqI: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 563 | "(\<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 | 564 | 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 | 565 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 566 | lemma mset_subset_eq_count: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 567 | "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 | 568 | by (simp add: subseteq_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 569 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 570 | lemma mset_subset_eq_exists_conv: "(A::'a multiset) \<subseteq># B \<longleftrightarrow> (\<exists>C. B = A + C)" | 
| 60678 | 571 | unfolding subseteq_mset_def | 
| 80095 | 572 | by (metis add_diff_cancel_left' count_diff count_union le_Suc_ex le_add_same_cancel1 multiset_eq_iff zero_le) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 573 | |
| 67398 | 574 | interpretation subset_mset: ordered_cancel_comm_monoid_diff "(+)" 0 "(\<subseteq>#)" "(\<subset>#)" "(-)" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 575 | by standard (simp, fact mset_subset_eq_exists_conv) | 
| 64585 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 haftmann parents: 
64531diff
changeset | 576 | \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 577 | |
| 64017 
6e7bf7678518
more multiset simp rules
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63924diff
changeset | 578 | declare subset_mset.add_diff_assoc[simp] subset_mset.add_diff_assoc2[simp] | 
| 
6e7bf7678518
more multiset simp rules
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63924diff
changeset | 579 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 580 | lemma mset_subset_eq_mono_add_right_cancel: "(A::'a multiset) + C \<subseteq># B + C \<longleftrightarrow> A \<subseteq># B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 581 | by (fact subset_mset.add_le_cancel_right) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 582 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 583 | lemma mset_subset_eq_mono_add_left_cancel: "C + (A::'a multiset) \<subseteq># C + B \<longleftrightarrow> A \<subseteq># B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 584 | by (fact subset_mset.add_le_cancel_left) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 585 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 586 | lemma mset_subset_eq_mono_add: "(A::'a multiset) \<subseteq># B \<Longrightarrow> C \<subseteq># D \<Longrightarrow> A + C \<subseteq># B + D" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 587 | by (fact subset_mset.add_mono) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 588 | |
| 63560 
3e3097ac37d1
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63547diff
changeset | 589 | lemma mset_subset_eq_add_left: "(A::'a multiset) \<subseteq># A + B" | 
| 
3e3097ac37d1
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63547diff
changeset | 590 | by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 591 | |
| 63560 
3e3097ac37d1
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63547diff
changeset | 592 | lemma mset_subset_eq_add_right: "B \<subseteq># (A::'a multiset) + B" | 
| 
3e3097ac37d1
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63547diff
changeset | 593 | by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 594 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 595 | lemma single_subset_iff [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 596 |   "{#a#} \<subseteq># M \<longleftrightarrow> a \<in># M"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 597 | 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 | 598 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 599 | lemma mset_subset_eq_single: "a \<in># B \<Longrightarrow> {#a#} \<subseteq># B"
 | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 600 | by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 601 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 602 | lemma mset_subset_eq_add_mset_cancel: \<open>add_mset a A \<subseteq># add_mset a B \<longleftrightarrow> A \<subseteq># B\<close> | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 603 | unfolding add_mset_add_single[of _ A] add_mset_add_single[of _ B] | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 604 | by (rule mset_subset_eq_mono_add_right_cancel) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 605 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 606 | lemma multiset_diff_union_assoc: | 
| 60606 | 607 | fixes A B C D :: "'a multiset" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 608 | 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 | 609 | by (fact subset_mset.diff_add_assoc) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 610 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 611 | lemma mset_subset_eq_multiset_union_diff_commute: | 
| 60606 | 612 | fixes A B C D :: "'a multiset" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 613 | 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 | 614 | by (fact subset_mset.add_diff_assoc2) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 615 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 616 | lemma diff_subset_eq_self[simp]: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 617 | "(M::'a multiset) - N \<subseteq># M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 618 | by (simp add: subseteq_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 619 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 620 | lemma mset_subset_eqD: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 621 | assumes "A \<subseteq># B" and "x \<in># A" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 622 | shows "x \<in># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 623 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 624 | 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 | 625 | 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 | 626 | by (simp add: subseteq_mset_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 627 | finally show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 628 | qed | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 629 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 630 | lemma mset_subsetD: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 631 | "A \<subset># B \<Longrightarrow> x \<in># A \<Longrightarrow> x \<in># B" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 632 | by (auto intro: mset_subset_eqD [of A]) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 633 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 634 | lemma set_mset_mono: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 635 | "A \<subseteq># B \<Longrightarrow> set_mset A \<subseteq> set_mset B" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 636 | by (metis mset_subset_eqD subsetI) | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 637 | |
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 638 | lemma mset_subset_eq_insertD: | 
| 80095 | 639 | assumes "add_mset x A \<subseteq># B" | 
| 640 | shows "x \<in># B \<and> A \<subset># B" | |
| 641 | proof | |
| 642 | show "x \<in># B" | |
| 643 | using assms by (simp add: mset_subset_eqD) | |
| 644 | have "A \<subseteq># add_mset x A" | |
| 645 | by (metis (no_types) add_mset_add_single mset_subset_eq_add_left) | |
| 646 | then have "A \<subset># add_mset x A" | |
| 647 | by (meson multi_self_add_other_not_self subset_mset.le_imp_less_or_eq) | |
| 648 | then show "A \<subset># B" | |
| 649 | using assms subset_mset.strict_trans2 by blast | |
| 650 | 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 | 651 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 652 | lemma mset_subset_insertD: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 653 | "add_mset x A \<subset># B \<Longrightarrow> x \<in># B \<and> A \<subset># B" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 654 | by (rule mset_subset_eq_insertD) simp | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 655 | |
| 63831 | 656 | lemma mset_subset_of_empty[simp]: "A \<subset># {#} \<longleftrightarrow> False"
 | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 657 | by (simp only: subset_mset.not_less_zero) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 658 | |
| 64587 | 659 | lemma empty_subset_add_mset[simp]: "{#} \<subset># add_mset x M"
 | 
| 660 | by (auto intro: subset_mset.gr_zeroI) | |
| 63831 | 661 | |
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 662 | lemma empty_le: "{#} \<subseteq># A"
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 663 | by (fact subset_mset.zero_le) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 664 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 665 | lemma insert_subset_eq_iff: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 666 |   "add_mset a A \<subseteq># B \<longleftrightarrow> a \<in># B \<and> A \<subseteq># B - {#a#}"
 | 
| 80095 | 667 | using mset_subset_eq_insertD subset_mset.le_diff_conv2 by fastforce | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 668 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 669 | lemma insert_union_subset_iff: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 670 |   "add_mset a A \<subset># B \<longleftrightarrow> a \<in># B \<and> A \<subset># B - {#a#}"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 671 | by (auto simp add: insert_subset_eq_iff subset_mset_def) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 672 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 673 | lemma subset_eq_diff_conv: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 674 | "A - C \<subseteq># B \<longleftrightarrow> A \<subseteq># B + C" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 675 | by (simp add: subseteq_mset_def le_diff_conv) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 676 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 677 | lemma multi_psub_of_add_self [simp]: "A \<subset># add_mset x A" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 678 | by (auto simp: subset_mset_def subseteq_mset_def) | 
| 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 679 | |
| 64076 | 680 | lemma multi_psub_self: "A \<subset># A = False" | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 681 | 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 | 682 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 683 | lemma mset_subset_add_mset [simp]: "add_mset x N \<subset># add_mset x M \<longleftrightarrow> N \<subset># M" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 684 | unfolding add_mset_add_single[of _ N] add_mset_add_single[of _ M] | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 685 | by (fact subset_mset.add_less_cancel_right) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 686 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 687 | lemma mset_subset_diff_self: "c \<in># B \<Longrightarrow> B - {#c#} \<subset># B"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 688 | by (auto simp: subset_mset_def elim: mset_add) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 689 | |
| 64077 | 690 | lemma Diff_eq_empty_iff_mset: "A - B = {#} \<longleftrightarrow> A \<subseteq># B"
 | 
| 691 | by (auto simp: multiset_eq_iff subseteq_mset_def) | |
| 692 | ||
| 64418 | 693 | lemma add_mset_subseteq_single_iff[iff]: "add_mset a M \<subseteq># {#b#} \<longleftrightarrow> M = {#} \<and> a = b"
 | 
| 694 | proof | |
| 695 |   assume A: "add_mset a M \<subseteq># {#b#}"
 | |
| 696 | then have \<open>a = b\<close> | |
| 697 | by (auto dest: mset_subset_eq_insertD) | |
| 698 |   then show "M={#} \<and> a=b"
 | |
| 699 | using A by (simp add: mset_subset_eq_add_mset_cancel) | |
| 700 | qed simp | |
| 701 | ||
| 79800 | 702 | lemma nonempty_subseteq_mset_eq_single: "M \<noteq> {#} \<Longrightarrow> M \<subseteq># {#x#} \<Longrightarrow> M = {#x#}"
 | 
| 703 | by (cases M) (metis single_is_union subset_mset.less_eqE) | |
| 704 | ||
| 705 | lemma nonempty_subseteq_mset_iff_single: "(M \<noteq> {#} \<and> M \<subseteq># {#x#} \<and> P) \<longleftrightarrow> M = {#x#} \<and> P"
 | |
| 706 | by (cases M) (metis empty_not_add_mset nonempty_subseteq_mset_eq_single subset_mset.order_refl) | |
| 707 | ||
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 708 | |
| 64076 | 709 | subsubsection \<open>Intersection and bounded union\<close> | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 710 | |
| 73393 | 711 | definition inter_mset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> (infixl \<open>\<inter>#\<close> 70) | 
| 712 | where \<open>A \<inter># B = A - (A - B)\<close> | |
| 713 | ||
| 73451 | 714 | lemma count_inter_mset [simp]: | 
| 715 | \<open>count (A \<inter># B) x = min (count A x) (count B x)\<close> | |
| 716 | by (simp add: inter_mset_def) | |
| 717 | ||
| 718 | (*global_interpretation subset_mset: semilattice_order \<open>(\<inter>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> | |
| 719 | by standard (simp_all add: multiset_eq_iff subseteq_mset_def subset_mset_def min_def)*) | |
| 720 | ||
| 73393 | 721 | interpretation subset_mset: semilattice_inf \<open>(\<inter>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> | 
| 73451 | 722 | by standard (simp_all add: multiset_eq_iff subseteq_mset_def) | 
| 723 | \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 724 | |
| 73393 | 725 | definition union_mset :: \<open>'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> (infixl \<open>\<union>#\<close> 70) | 
| 726 | where \<open>A \<union># B = A + (B - A)\<close> | |
| 727 | ||
| 73451 | 728 | lemma count_union_mset [simp]: | 
| 729 | \<open>count (A \<union># B) x = max (count A x) (count B x)\<close> | |
| 730 | by (simp add: union_mset_def) | |
| 731 | ||
| 732 | global_interpretation subset_mset: semilattice_neutr_order \<open>(\<union>#)\<close> \<open>{#}\<close> \<open>(\<supseteq>#)\<close> \<open>(\<supset>#)\<close>
 | |
| 80095 | 733 | proof | 
| 734 | show "\<And>a b. (b \<subseteq># a) = (a = a \<union># b)" | |
| 735 | by (simp add: Diff_eq_empty_iff_mset union_mset_def) | |
| 736 | show "\<And>a b. (b \<subset># a) = (a = a \<union># b \<and> a \<noteq> b)" | |
| 737 | by (metis Diff_eq_empty_iff_mset add_cancel_left_right subset_mset_def union_mset_def) | |
| 738 | qed (auto simp: multiset_eqI union_mset_def) | |
| 73451 | 739 | |
| 73393 | 740 | interpretation subset_mset: semilattice_sup \<open>(\<union>#)\<close> \<open>(\<subseteq>#)\<close> \<open>(\<subset>#)\<close> | 
| 64076 | 741 | proof - | 
| 742 | have [simp]: "m \<le> n \<Longrightarrow> q \<le> n \<Longrightarrow> m + (q - m) \<le> n" for m n q :: nat | |
| 743 | by arith | |
| 67398 | 744 | show "class.semilattice_sup (\<union>#) (\<subseteq>#) (\<subset>#)" | 
| 73393 | 745 | by standard (auto simp add: union_mset_def subseteq_mset_def) | 
| 64585 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 haftmann parents: 
64531diff
changeset | 746 | qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 64076 | 747 | |
| 67398 | 748 | interpretation subset_mset: bounded_lattice_bot "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)" | 
| 749 |   "(\<union>#)" "{#}"
 | |
| 64076 | 750 | by standard auto | 
| 64585 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 haftmann parents: 
64531diff
changeset | 751 | \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 64076 | 752 | |
| 753 | ||
| 754 | subsubsection \<open>Additional intersection facts\<close> | |
| 755 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 756 | lemma set_mset_inter [simp]: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 757 | "set_mset (A \<inter># B) = set_mset A \<inter> set_mset B" | 
| 73393 | 758 | by (simp only: set_mset_def) auto | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 759 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 760 | lemma diff_intersect_left_idem [simp]: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 761 | "M - M \<inter># N = M - N" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 762 | by (simp add: multiset_eq_iff min_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 763 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 764 | lemma diff_intersect_right_idem [simp]: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 765 | "M - N \<inter># M = M - N" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 766 | by (simp add: multiset_eq_iff min_def) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 767 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 768 | lemma multiset_inter_single[simp]: "a \<noteq> b \<Longrightarrow> {#a#} \<inter># {#b#} = {#}"
 | 
| 46730 | 769 | 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 | 770 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 771 | lemma multiset_union_diff_commute: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 772 |   assumes "B \<inter># C = {#}"
 | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 773 | 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 | 774 | proof (rule multiset_eqI) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 775 | fix x | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 776 | from assms have "min (count B x) (count C x) = 0" | 
| 46730 | 777 | by (auto simp add: multiset_eq_iff) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 778 | 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 | 779 | unfolding min_def by (auto split: if_splits) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 780 | then show "count (A + B - C) x = count (A - C + B) x" | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 781 | by auto | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 782 | qed | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 783 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 784 | lemma disjunct_not_in: | 
| 80095 | 785 |   "A \<inter># B = {#} \<longleftrightarrow> (\<forall>a. a \<notin># A \<or> a \<notin># B)"
 | 
| 786 | by (metis disjoint_iff set_mset_eq_empty_iff set_mset_inter) | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 787 | |
| 64077 | 788 | lemma inter_mset_empty_distrib_right: "A \<inter># (B + C) = {#} \<longleftrightarrow> A \<inter># B = {#} \<and> A \<inter># C = {#}"
 | 
| 789 | by (meson disjunct_not_in union_iff) | |
| 790 | ||
| 791 | lemma inter_mset_empty_distrib_left: "(A + B) \<inter># C = {#} \<longleftrightarrow> A \<inter># C = {#} \<and> B \<inter># C = {#}"
 | |
| 792 | by (meson disjunct_not_in union_iff) | |
| 793 | ||
| 73393 | 794 | lemma add_mset_inter_add_mset [simp]: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 795 | "add_mset a A \<inter># add_mset a B = add_mset a (A \<inter># B)" | 
| 73393 | 796 | by (rule multiset_eqI) simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 797 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 798 | lemma add_mset_disjoint [simp]: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 799 |   "add_mset a A \<inter># B = {#} \<longleftrightarrow> a \<notin># B \<and> A \<inter># B = {#}"
 | 
| 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 800 |   "{#} = add_mset a A \<inter># B \<longleftrightarrow> a \<notin># B \<and> {#} = A \<inter># B"
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 801 | by (auto simp: disjunct_not_in) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 802 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 803 | lemma disjoint_add_mset [simp]: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 804 |   "B \<inter># add_mset a A = {#} \<longleftrightarrow> a \<notin># B \<and> B \<inter># A = {#}"
 | 
| 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 805 |   "{#} = A \<inter># add_mset b B \<longleftrightarrow> b \<notin># A \<and> {#} = A \<inter># B"
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 806 | by (auto simp: disjunct_not_in) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 807 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 808 | lemma inter_add_left1: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) \<inter># N = M \<inter># N" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 809 | 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 | 810 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 811 | lemma inter_add_left2: "x \<in># N \<Longrightarrow> (add_mset x M) \<inter># N = add_mset x (M \<inter># (N - {#x#}))"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 812 | 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 | 813 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 814 | lemma inter_add_right1: "\<not> x \<in># N \<Longrightarrow> N \<inter># (add_mset x M) = N \<inter># M" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 815 | 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 | 816 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 817 | lemma inter_add_right2: "x \<in># N \<Longrightarrow> N \<inter># (add_mset x M) = add_mset x ((N - {#x#}) \<inter># M)"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 818 | 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 | 819 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 820 | lemma disjunct_set_mset_diff: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 821 |   assumes "M \<inter># N = {#}"
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 822 | shows "set_mset (M - N) = set_mset M" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 823 | proof (rule set_eqI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 824 | fix a | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 825 | 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 | 826 | by (simp add: disjunct_not_in) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 827 | 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 | 828 | 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 | 829 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 830 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 831 | lemma at_most_one_mset_mset_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 832 |   assumes "a \<notin># M - {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 833 |   shows "set_mset (M - {#a#}) = set_mset M - {a}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 834 | 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 | 835 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 836 | lemma more_than_one_mset_mset_diff: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 837 |   assumes "a \<in># M - {#a#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 838 |   shows "set_mset (M - {#a#}) = set_mset M"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 839 | proof (rule set_eqI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 840 | fix b | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 841 | 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 | 842 |   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 | 843 | using assms by (auto simp add: in_diff_count) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 844 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 845 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 846 | lemma inter_iff: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 847 | "a \<in># A \<inter># B \<longleftrightarrow> a \<in># A \<and> a \<in># B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 848 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 849 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 850 | lemma inter_union_distrib_left: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 851 | "A \<inter># B + C = (A + C) \<inter># (B + C)" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 852 | 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 | 853 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 854 | lemma inter_union_distrib_right: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 855 | "C + A \<inter># B = (C + A) \<inter># (C + B)" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 856 | 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 | 857 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 858 | lemma inter_subset_eq_union: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 859 | "A \<inter># B \<subseteq># A + B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 860 | 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 | 861 | |
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 862 | |
| 64076 | 863 | subsubsection \<open>Additional bounded union facts\<close> | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 864 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 865 | lemma set_mset_sup [simp]: | 
| 73393 | 866 | \<open>set_mset (A \<union># B) = set_mset A \<union> set_mset B\<close> | 
| 867 | by (simp only: set_mset_def) (auto simp add: less_max_iff_disj) | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 868 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 869 | lemma sup_union_left1 [simp]: "\<not> x \<in># N \<Longrightarrow> (add_mset x M) \<union># N = add_mset x (M \<union># N)" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 870 | by (simp add: multiset_eq_iff not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 871 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 872 | lemma sup_union_left2: "x \<in># N \<Longrightarrow> (add_mset x M) \<union># N = add_mset x (M \<union># (N - {#x#}))"
 | 
| 51623 | 873 | by (simp add: multiset_eq_iff) | 
| 874 | ||
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 875 | lemma sup_union_right1 [simp]: "\<not> x \<in># N \<Longrightarrow> N \<union># (add_mset x M) = add_mset x (N \<union># M)" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 876 | by (simp add: multiset_eq_iff not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 877 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 878 | lemma sup_union_right2: "x \<in># N \<Longrightarrow> N \<union># (add_mset x M) = add_mset x ((N - {#x#}) \<union># M)"
 | 
| 51623 | 879 | by (simp add: multiset_eq_iff) | 
| 880 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 881 | lemma sup_union_distrib_left: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 882 | "A \<union># B + C = (A + C) \<union># (B + C)" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 883 | 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 | 884 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 885 | lemma union_sup_distrib_right: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 886 | "C + A \<union># B = (C + A) \<union># (C + B)" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 887 | 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 | 888 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 889 | lemma union_diff_inter_eq_sup: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 890 | "A + B - A \<inter># B = A \<union># B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 891 | by (auto simp add: multiset_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 892 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 893 | lemma union_diff_sup_eq_inter: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 894 | "A + B - A \<union># B = A \<inter># B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 895 | by (auto simp add: multiset_eq_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 896 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 897 | lemma add_mset_union: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 898 | \<open>add_mset a A \<union># add_mset a B = add_mset a (A \<union># B)\<close> | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 899 | by (auto simp: multiset_eq_iff max_def) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 900 | |
| 51623 | 901 | |
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 902 | subsection \<open>Replicate and repeat operations\<close> | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 903 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 904 | definition replicate_mset :: "nat \<Rightarrow> 'a \<Rightarrow> 'a multiset" where | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 905 |   "replicate_mset n x = (add_mset x ^^ n) {#}"
 | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 906 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 907 | lemma replicate_mset_0[simp]: "replicate_mset 0 x = {#}"
 | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 908 | unfolding replicate_mset_def by simp | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 909 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 910 | lemma replicate_mset_Suc [simp]: "replicate_mset (Suc n) x = add_mset x (replicate_mset n x)" | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 911 | unfolding replicate_mset_def by (induct n) (auto intro: add.commute) | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 912 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 913 | lemma count_replicate_mset[simp]: "count (replicate_mset n x) y = (if y = x then n else 0)" | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 914 | unfolding replicate_mset_def by (induct n) auto | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 915 | |
| 73393 | 916 | lift_definition repeat_mset :: \<open>nat \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset\<close> | 
| 917 | is \<open>\<lambda>n M a. n * M a\<close> by simp | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 918 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 919 | lemma count_repeat_mset [simp]: "count (repeat_mset i A) a = i * count A a" | 
| 73393 | 920 | by transfer rule | 
| 921 | ||
| 922 | lemma repeat_mset_0 [simp]: | |
| 923 |   \<open>repeat_mset 0 M = {#}\<close>
 | |
| 924 | by transfer simp | |
| 925 | ||
| 926 | lemma repeat_mset_Suc [simp]: | |
| 927 | \<open>repeat_mset (Suc n) M = M + repeat_mset n M\<close> | |
| 928 | by transfer simp | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 929 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 930 | lemma repeat_mset_right [simp]: "repeat_mset a (repeat_mset b A) = repeat_mset (a * b) A" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 931 | by (auto simp: multiset_eq_iff left_diff_distrib') | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 932 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 933 | lemma left_diff_repeat_mset_distrib': \<open>repeat_mset (i - j) u = repeat_mset i u - repeat_mset j u\<close> | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 934 | by (auto simp: multiset_eq_iff left_diff_distrib') | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 935 | |
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 936 | lemma left_add_mult_distrib_mset: | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 937 | "repeat_mset i u + (repeat_mset j u + k) = repeat_mset (i+j) u + k" | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 938 | by (auto simp: multiset_eq_iff add_mult_distrib) | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 939 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 940 | lemma repeat_mset_distrib: | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 941 | "repeat_mset (m + n) A = repeat_mset m A + repeat_mset n A" | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 942 | by (auto simp: multiset_eq_iff Nat.add_mult_distrib) | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 943 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 944 | lemma repeat_mset_distrib2[simp]: | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 945 | "repeat_mset n (A + B) = repeat_mset n A + repeat_mset n B" | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 946 | by (auto simp: multiset_eq_iff add_mult_distrib2) | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 947 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 948 | lemma repeat_mset_replicate_mset[simp]: | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 949 |   "repeat_mset n {#a#} = replicate_mset n a"
 | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 950 | by (auto simp: multiset_eq_iff) | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 951 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 952 | lemma repeat_mset_distrib_add_mset[simp]: | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 953 | "repeat_mset n (add_mset a A) = replicate_mset n a + repeat_mset n A" | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 954 | by (auto simp: multiset_eq_iff) | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 955 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 956 | lemma repeat_mset_empty[simp]: "repeat_mset n {#} = {#}"
 | 
| 73393 | 957 | by transfer simp | 
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 958 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 959 | |
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 960 | subsubsection \<open>Simprocs\<close> | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 961 | |
| 65031 
52e2c99f3711
use the cancellation simprocs directly
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65029diff
changeset | 962 | lemma repeat_mset_iterate_add: \<open>repeat_mset n M = iterate_add n M\<close> | 
| 
52e2c99f3711
use the cancellation simprocs directly
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65029diff
changeset | 963 | unfolding iterate_add_def by (induction n) auto | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 964 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 965 | lemma mset_subseteq_add_iff1: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 966 | "j \<le> (i::nat) \<Longrightarrow> (repeat_mset i u + m \<subseteq># repeat_mset j u + n) = (repeat_mset (i-j) u + m \<subseteq># n)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 967 | by (auto simp add: subseteq_mset_def nat_le_add_iff1) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 968 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 969 | lemma mset_subseteq_add_iff2: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 970 | "i \<le> (j::nat) \<Longrightarrow> (repeat_mset i u + m \<subseteq># repeat_mset j u + n) = (m \<subseteq># repeat_mset (j-i) u + n)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 971 | by (auto simp add: subseteq_mset_def nat_le_add_iff2) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 972 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 973 | lemma mset_subset_add_iff1: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 974 | "j \<le> (i::nat) \<Longrightarrow> (repeat_mset i u + m \<subset># repeat_mset j u + n) = (repeat_mset (i-j) u + m \<subset># n)" | 
| 65031 
52e2c99f3711
use the cancellation simprocs directly
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65029diff
changeset | 975 | unfolding subset_mset_def repeat_mset_iterate_add | 
| 
52e2c99f3711
use the cancellation simprocs directly
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65029diff
changeset | 976 | by (simp add: iterate_add_eq_add_iff1 mset_subseteq_add_iff1[unfolded repeat_mset_iterate_add]) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 977 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 978 | lemma mset_subset_add_iff2: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 979 | "i \<le> (j::nat) \<Longrightarrow> (repeat_mset i u + m \<subset># repeat_mset j u + n) = (m \<subset># repeat_mset (j-i) u + n)" | 
| 65031 
52e2c99f3711
use the cancellation simprocs directly
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65029diff
changeset | 980 | unfolding subset_mset_def repeat_mset_iterate_add | 
| 
52e2c99f3711
use the cancellation simprocs directly
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65029diff
changeset | 981 | by (simp add: iterate_add_eq_add_iff2 mset_subseteq_add_iff2[unfolded repeat_mset_iterate_add]) | 
| 65029 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 982 | |
| 69605 | 983 | ML_file \<open>multiset_simprocs.ML\<close> | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 984 | |
| 65029 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 985 | lemma add_mset_replicate_mset_safe[cancelation_simproc_pre]: \<open>NO_MATCH {#} M \<Longrightarrow> add_mset a M = {#a#} + M\<close>
 | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 986 | by simp | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 987 | |
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 988 | declare repeat_mset_iterate_add[cancelation_simproc_pre] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 989 | |
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 990 | declare iterate_add_distrib[cancelation_simproc_pre] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 991 | declare repeat_mset_iterate_add[symmetric, cancelation_simproc_post] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 992 | |
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 993 | declare add_mset_not_empty[cancelation_simproc_eq_elim] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 994 | empty_not_add_mset[cancelation_simproc_eq_elim] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 995 | subset_mset.le_zero_eq[cancelation_simproc_eq_elim] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 996 | empty_not_add_mset[cancelation_simproc_eq_elim] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 997 | add_mset_not_empty[cancelation_simproc_eq_elim] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 998 | subset_mset.le_zero_eq[cancelation_simproc_eq_elim] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 999 | le_zero_eq[cancelation_simproc_eq_elim] | 
| 
00731700e54f
cancellation simprocs generalising the multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
65027diff
changeset | 1000 | |
| 65027 
2b8583507891
renaming multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
64911diff
changeset | 1001 | simproc_setup mseteq_cancel | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1002 |   ("(l::'a multiset) + m = n" | "(l::'a multiset) = m + n" |
 | 
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1003 | "add_mset a m = n" | "m = add_mset a n" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1004 | "replicate_mset p a = n" | "m = replicate_mset p a" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1005 | "repeat_mset p m = n" | "m = repeat_mset p m") = | 
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
77987diff
changeset | 1006 | \<open>K Cancel_Simprocs.eq_cancel\<close> | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1007 | |
| 65027 
2b8583507891
renaming multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
64911diff
changeset | 1008 | simproc_setup msetsubset_cancel | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1009 |   ("(l::'a multiset) + m \<subset># n" | "(l::'a multiset) \<subset># m + n" |
 | 
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1010 | "add_mset a m \<subset># n" | "m \<subset># add_mset a n" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1011 | "replicate_mset p r \<subset># n" | "m \<subset># replicate_mset p r" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1012 | "repeat_mset p m \<subset># n" | "m \<subset># repeat_mset p m") = | 
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
77987diff
changeset | 1013 | \<open>K Multiset_Simprocs.subset_cancel_msets\<close> | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1014 | |
| 65027 
2b8583507891
renaming multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
64911diff
changeset | 1015 | simproc_setup msetsubset_eq_cancel | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1016 |   ("(l::'a multiset) + m \<subseteq># n" | "(l::'a multiset) \<subseteq># m + n" |
 | 
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1017 | "add_mset a m \<subseteq># n" | "m \<subseteq># add_mset a n" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1018 | "replicate_mset p r \<subseteq># n" | "m \<subseteq># replicate_mset p r" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1019 | "repeat_mset p m \<subseteq># n" | "m \<subseteq># repeat_mset p m") = | 
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
77987diff
changeset | 1020 | \<open>K Multiset_Simprocs.subseteq_cancel_msets\<close> | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1021 | |
| 65027 
2b8583507891
renaming multiset simprocs
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
64911diff
changeset | 1022 | simproc_setup msetdiff_cancel | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1023 |   ("((l::'a multiset) + m) - n" | "(l::'a multiset) - (m + n)" |
 | 
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1024 | "add_mset a m - n" | "m - add_mset a n" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1025 | "replicate_mset p r - n" | "m - replicate_mset p r" | | 
| 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 1026 | "repeat_mset p m - n" | "m - repeat_mset p m") = | 
| 78099 
4d9349989d94
more uniform simproc_setup: avoid vacuous abstraction over morphism, which sometimes captures context values in its functional closure;
 wenzelm parents: 
77987diff
changeset | 1027 | \<open>K Cancel_Simprocs.diff_cancel\<close> | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1028 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1029 | |
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1030 | subsubsection \<open>Conditionally complete lattice\<close> | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1031 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1032 | instantiation multiset :: (type) Inf | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1033 | begin | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1034 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1035 | lift_definition Inf_multiset :: "'a multiset set \<Rightarrow> 'a multiset" is | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1036 |   "\<lambda>A i. if A = {} then 0 else Inf ((\<lambda>f. f i) ` A)"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1037 | proof - | 
| 73270 | 1038 |   fix A :: "('a \<Rightarrow> nat) set"
 | 
| 1039 |   assume *: "\<And>f. f \<in> A \<Longrightarrow> finite {x. 0 < f x}"
 | |
| 1040 |   show \<open>finite {i. 0 < (if A = {} then 0 else INF f\<in>A. f i)}\<close>
 | |
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1041 |   proof (cases "A = {}")
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1042 | case False | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1043 | then obtain f where "f \<in> A" by blast | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1044 |     hence "{i. Inf ((\<lambda>f. f i) ` A) > 0} \<subseteq> {i. f i > 0}"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1045 | by (auto intro: less_le_trans[OF _ cInf_lower]) | 
| 73270 | 1046 | moreover from \<open>f \<in> A\<close> * have "finite \<dots>" by simp | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1047 |     ultimately have "finite {i. Inf ((\<lambda>f. f i) ` A) > 0}" by (rule finite_subset)
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1048 | with False show ?thesis by simp | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1049 | qed simp_all | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1050 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1051 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1052 | instance .. | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1053 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1054 | end | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1055 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1056 | lemma Inf_multiset_empty: "Inf {} = {#}"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1057 | by transfer simp_all | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1058 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1059 | lemma count_Inf_multiset_nonempty: "A \<noteq> {} \<Longrightarrow> count (Inf A) x = Inf ((\<lambda>X. count X x) ` A)"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1060 | by transfer simp_all | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1061 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1062 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1063 | instantiation multiset :: (type) Sup | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1064 | begin | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1065 | |
| 63360 | 1066 | definition Sup_multiset :: "'a multiset set \<Rightarrow> 'a multiset" where | 
| 1067 |   "Sup_multiset A = (if A \<noteq> {} \<and> subset_mset.bdd_above A then
 | |
| 1068 |            Abs_multiset (\<lambda>i. Sup ((\<lambda>X. count X i) ` A)) else {#})"
 | |
| 1069 | ||
| 1070 | lemma Sup_multiset_empty: "Sup {} = {#}"
 | |
| 1071 | by (simp add: Sup_multiset_def) | |
| 1072 | ||
| 73451 | 1073 | lemma Sup_multiset_unbounded: "\<not> subset_mset.bdd_above A \<Longrightarrow> Sup A = {#}"
 | 
| 63360 | 1074 | by (simp add: Sup_multiset_def) | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1075 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1076 | instance .. | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1077 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1078 | end | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1079 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1080 | lemma bdd_above_multiset_imp_bdd_above_count: | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1081 | assumes "subset_mset.bdd_above (A :: 'a multiset set)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1082 | shows "bdd_above ((\<lambda>X. count X x) ` A)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1083 | proof - | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1084 | from assms obtain Y where Y: "\<forall>X\<in>A. X \<subseteq># Y" | 
| 73451 | 1085 | by (meson subset_mset.bdd_above.E) | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1086 | hence "count X x \<le> count Y x" if "X \<in> A" for X | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1087 | using that by (auto intro: mset_subset_eq_count) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1088 | thus ?thesis by (intro bdd_aboveI[of _ "count Y x"]) auto | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1089 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1090 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1091 | lemma bdd_above_multiset_imp_finite_support: | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1092 |   assumes "A \<noteq> {}" "subset_mset.bdd_above (A :: 'a multiset set)"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1093 |   shows   "finite (\<Union>X\<in>A. {x. count X x > 0})"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1094 | proof - | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1095 | from assms obtain Y where Y: "\<forall>X\<in>A. X \<subseteq># Y" | 
| 73451 | 1096 | by (meson subset_mset.bdd_above.E) | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1097 | hence "count X x \<le> count Y x" if "X \<in> A" for X x | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1098 | using that by (auto intro: mset_subset_eq_count) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1099 |   hence "(\<Union>X\<in>A. {x. count X x > 0}) \<subseteq> {x. count Y x > 0}"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1100 | by safe (erule less_le_trans) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1101 | moreover have "finite \<dots>" by simp | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1102 | ultimately show ?thesis by (rule finite_subset) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1103 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1104 | |
| 63360 | 1105 | lemma Sup_multiset_in_multiset: | 
| 73270 | 1106 |   \<open>finite {i. 0 < (SUP M\<in>A. count M i)}\<close>
 | 
| 1107 |   if \<open>A \<noteq> {}\<close> \<open>subset_mset.bdd_above A\<close>
 | |
| 1108 | proof - | |
| 63360 | 1109 |   have "{i. Sup ((\<lambda>X. count X i) ` A) > 0} \<subseteq> (\<Union>X\<in>A. {i. 0 < count X i})"
 | 
| 1110 | proof safe | |
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69107diff
changeset | 1111 | fix i assume pos: "(SUP X\<in>A. count X i) > 0" | 
| 63360 | 1112 |     show "i \<in> (\<Union>X\<in>A. {i. 0 < count X i})"
 | 
| 1113 | proof (rule ccontr) | |
| 1114 |       assume "i \<notin> (\<Union>X\<in>A. {i. 0 < count X i})"
 | |
| 1115 | hence "\<forall>X\<in>A. count X i \<le> 0" by (auto simp: count_eq_zero_iff) | |
| 73270 | 1116 | with that have "(SUP X\<in>A. count X i) \<le> 0" | 
| 63360 | 1117 | by (intro cSup_least bdd_above_multiset_imp_bdd_above_count) auto | 
| 1118 | with pos show False by simp | |
| 1119 | qed | |
| 1120 | qed | |
| 73270 | 1121 | moreover from that have "finite \<dots>" | 
| 1122 | by (rule bdd_above_multiset_imp_finite_support) | |
| 1123 |   ultimately show "finite {i. Sup ((\<lambda>X. count X i) ` A) > 0}"
 | |
| 1124 | by (rule finite_subset) | |
| 63360 | 1125 | qed | 
| 1126 | ||
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1127 | lemma count_Sup_multiset_nonempty: | 
| 73270 | 1128 | \<open>count (Sup A) x = (SUP X\<in>A. count X x)\<close> | 
| 1129 |   if \<open>A \<noteq> {}\<close> \<open>subset_mset.bdd_above A\<close>
 | |
| 1130 | using that by (simp add: Sup_multiset_def Sup_multiset_in_multiset count_Abs_multiset) | |
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1131 | |
| 67398 | 1132 | interpretation subset_mset: conditionally_complete_lattice Inf Sup "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)" "(\<union>#)" | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1133 | proof | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1134 | fix X :: "'a multiset" and A | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1135 | assume "X \<in> A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1136 | show "Inf A \<subseteq># X" | 
| 80095 | 1137 | by (metis \<open>X \<in> A\<close> count_Inf_multiset_nonempty empty_iff image_eqI mset_subset_eqI wellorder_Inf_le1) | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1138 | next | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1139 | fix X :: "'a multiset" and A | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1140 |   assume nonempty: "A \<noteq> {}" and le: "\<And>Y. Y \<in> A \<Longrightarrow> X \<subseteq># Y"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1141 | show "X \<subseteq># Inf A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1142 | proof (rule mset_subset_eqI) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1143 | fix x | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69107diff
changeset | 1144 | from nonempty have "count X x \<le> (INF X\<in>A. count X x)" | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1145 | by (intro cInf_greatest) (auto intro: mset_subset_eq_count le) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1146 | also from nonempty have "\<dots> = count (Inf A) x" by (simp add: count_Inf_multiset_nonempty) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1147 | finally show "count X x \<le> count (Inf A) x" . | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1148 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1149 | next | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1150 | fix X :: "'a multiset" and A | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1151 | assume X: "X \<in> A" and bdd: "subset_mset.bdd_above A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1152 | show "X \<subseteq># Sup A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1153 | proof (rule mset_subset_eqI) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1154 | fix x | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1155 |     from X have "A \<noteq> {}" by auto
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69107diff
changeset | 1156 | have "count X x \<le> (SUP X\<in>A. count X x)" | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1157 | by (intro cSUP_upper X bdd_above_multiset_imp_bdd_above_count bdd) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1158 |     also from count_Sup_multiset_nonempty[OF \<open>A \<noteq> {}\<close> bdd]
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69107diff
changeset | 1159 | have "(SUP X\<in>A. count X x) = count (Sup A) x" by simp | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1160 | finally show "count X x \<le> count (Sup A) x" . | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1161 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1162 | next | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1163 | fix X :: "'a multiset" and A | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1164 |   assume nonempty: "A \<noteq> {}" and ge: "\<And>Y. Y \<in> A \<Longrightarrow> Y \<subseteq># X"
 | 
| 73451 | 1165 | from ge have bdd: "subset_mset.bdd_above A" | 
| 1166 | by blast | |
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1167 | show "Sup A \<subseteq># X" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1168 | proof (rule mset_subset_eqI) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1169 | fix x | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1170 |     from count_Sup_multiset_nonempty[OF \<open>A \<noteq> {}\<close> bdd]
 | 
| 69260 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 haftmann parents: 
69107diff
changeset | 1171 | have "count (Sup A) x = (SUP X\<in>A. count X x)" . | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1172 | also from nonempty have "\<dots> \<le> count X x" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1173 | by (intro cSup_least) (auto intro: mset_subset_eq_count ge) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1174 | finally show "count (Sup A) x \<le> count X x" . | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1175 | qed | 
| 64585 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 haftmann parents: 
64531diff
changeset | 1176 | qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1177 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1178 | lemma set_mset_Inf: | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1179 |   assumes "A \<noteq> {}"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1180 | shows "set_mset (Inf A) = (\<Inter>X\<in>A. set_mset X)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1181 | proof safe | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1182 | fix x X assume "x \<in># Inf A" "X \<in> A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1183 |   hence nonempty: "A \<noteq> {}" by (auto simp: Inf_multiset_empty)
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1184 |   from \<open>x \<in># Inf A\<close> have "{#x#} \<subseteq># Inf A" by auto
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1185 | also from \<open>X \<in> A\<close> have "\<dots> \<subseteq># X" by (rule subset_mset.cInf_lower) simp_all | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1186 | finally show "x \<in># X" by simp | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1187 | next | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1188 | fix x assume x: "x \<in> (\<Inter>X\<in>A. set_mset X)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1189 |   hence "{#x#} \<subseteq># X" if "X \<in> A" for X using that by auto
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1190 |   from assms and this have "{#x#} \<subseteq># Inf A" by (rule subset_mset.cInf_greatest)
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1191 | thus "x \<in># Inf A" by simp | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1192 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1193 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1194 | lemma in_Inf_multiset_iff: | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1195 |   assumes "A \<noteq> {}"
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1196 | shows "x \<in># Inf A \<longleftrightarrow> (\<forall>X\<in>A. x \<in># X)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1197 | proof - | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1198 | from assms have "set_mset (Inf A) = (\<Inter>X\<in>A. set_mset X)" by (rule set_mset_Inf) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1199 | also have "x \<in> \<dots> \<longleftrightarrow> (\<forall>X\<in>A. x \<in># X)" by simp | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1200 | finally show ?thesis . | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1201 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1202 | |
| 63360 | 1203 | lemma in_Inf_multisetD: "x \<in># Inf A \<Longrightarrow> X \<in> A \<Longrightarrow> x \<in># X" | 
| 1204 | by (subst (asm) in_Inf_multiset_iff) auto | |
| 1205 | ||
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1206 | lemma set_mset_Sup: | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1207 | assumes "subset_mset.bdd_above A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1208 | shows "set_mset (Sup A) = (\<Union>X\<in>A. set_mset X)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1209 | proof safe | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1210 | fix x assume "x \<in># Sup A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1211 |   hence nonempty: "A \<noteq> {}" by (auto simp: Sup_multiset_empty)
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1212 | show "x \<in> (\<Union>X\<in>A. set_mset X)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1213 | proof (rule ccontr) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1214 | assume x: "x \<notin> (\<Union>X\<in>A. set_mset X)" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1215 | have "count X x \<le> count (Sup A) x" if "X \<in> A" for X x | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1216 | using that by (intro mset_subset_eq_count subset_mset.cSup_upper assms) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1217 |     with x have "X \<subseteq># Sup A - {#x#}" if "X \<in> A" for X
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1218 | using that by (auto simp: subseteq_mset_def algebra_simps not_in_iff) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1219 |     hence "Sup A \<subseteq># Sup A - {#x#}" by (intro subset_mset.cSup_least nonempty)
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1220 | with \<open>x \<in># Sup A\<close> show False | 
| 80095 | 1221 | using mset_subset_diff_self by fastforce | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1222 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1223 | next | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1224 | fix x X assume "x \<in> set_mset X" "X \<in> A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1225 |   hence "{#x#} \<subseteq># X" by auto
 | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1226 | also have "X \<subseteq># Sup A" by (intro subset_mset.cSup_upper \<open>X \<in> A\<close> assms) | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1227 | finally show "x \<in> set_mset (Sup A)" by simp | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1228 | qed | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1229 | |
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1230 | lemma in_Sup_multiset_iff: | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1231 | assumes "subset_mset.bdd_above A" | 
| 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1232 | shows "x \<in># Sup A \<longleftrightarrow> (\<exists>X\<in>A. x \<in># X)" | 
| 80095 | 1233 | by (simp add: assms set_mset_Sup) | 
| 63358 
a500677d4cec
Conditionally complete lattice of multisets
 Manuel Eberl <eberlm@in.tum.de> parents: 
63310diff
changeset | 1234 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1235 | lemma in_Sup_multisetD: | 
| 63360 | 1236 | assumes "x \<in># Sup A" | 
| 1237 | shows "\<exists>X\<in>A. x \<in># X" | |
| 80095 | 1238 | using Sup_multiset_unbounded assms in_Sup_multiset_iff by fastforce | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63524diff
changeset | 1239 | |
| 67398 | 1240 | interpretation subset_mset: distrib_lattice "(\<inter>#)" "(\<subseteq>#)" "(\<subset>#)" "(\<union>#)" | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63524diff
changeset | 1241 | proof | 
| 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63524diff
changeset | 1242 | fix A B C :: "'a multiset" | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 1243 | show "A \<union># (B \<inter># C) = A \<union># B \<inter># (A \<union># C)" | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63524diff
changeset | 1244 | by (intro multiset_eqI) simp_all | 
| 64585 
2155c0c1ecb6
renewed and spread FIXME tags on watering bin interpretation, which got partially lost in 9f089287687b
 haftmann parents: 
64531diff
changeset | 1245 | qed \<comment> \<open>FIXME: avoid junk stemming from type class interpretation\<close> | 
| 63360 | 1246 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1247 | |
| 60500 | 1248 | subsubsection \<open>Filter (with comprehension syntax)\<close> | 
| 1249 | ||
| 1250 | text \<open>Multiset comprehension\<close> | |
| 41069 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 1251 | |
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1252 | 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 | 1253 | 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 | 1254 | by (rule filter_preserves_multiset) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1255 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1256 | syntax (ASCII) | 
| 63689 | 1257 |   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{#_ :# _./ _#})")
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1258 | syntax | 
| 63689 | 1259 |   "_MCollect" :: "pttrn \<Rightarrow> 'a multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"    ("(1{#_ \<in># _./ _#})")
 | 
| 80768 | 1260 | syntax_consts | 
| 1261 | "_MCollect" == filter_mset | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1262 | translations | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1263 |   "{#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 | 1264 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1265 | lemma count_filter_mset [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1266 | "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 | 1267 | by (simp add: filter_mset.rep_eq) | 
| 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1268 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1269 | lemma set_mset_filter [simp]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1270 |   "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 | 1271 | 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 | 1272 | |
| 60606 | 1273 | lemma filter_empty_mset [simp]: "filter_mset P {#} = {#}"
 | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1274 | by (rule multiset_eqI) simp | 
| 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1275 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1276 | lemma filter_single_mset: "filter_mset P {#x#} = (if P x then {#x#} else {#})"
 | 
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 1277 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1278 | |
| 60606 | 1279 | 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 | 1280 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 1281 | |
| 60606 | 1282 | 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 | 1283 | by (rule multiset_eqI) simp | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 1284 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 1285 | 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 | 1286 | by (rule multiset_eqI) simp | 
| 
6fabc0414055
name filter operation just filter (c.f. List.filter and list comprehension syntax)
 haftmann parents: 
40968diff
changeset | 1287 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 1288 | lemma filter_sup_mset[simp]: "filter_mset P (A \<union># B) = filter_mset P A \<union># filter_mset P B" | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 1289 | by (rule multiset_eqI) simp | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 1290 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1291 | lemma filter_mset_add_mset [simp]: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1292 | "filter_mset P (add_mset x A) = | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 1293 | (if P x then add_mset x (filter_mset P A) else filter_mset P A)" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1294 | by (auto simp: multiset_eq_iff) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1295 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1296 | lemma multiset_filter_subset[simp]: "filter_mset f M \<subseteq># M" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 1297 | by (simp add: mset_subset_eqI) | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 1298 | |
| 60606 | 1299 | lemma multiset_filter_mono: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1300 | assumes "A \<subseteq># B" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1301 | shows "filter_mset f A \<subseteq># filter_mset f B" | 
| 80095 | 1302 | by (metis assms filter_sup_mset subset_mset.order_iff) | 
| 58035 | 1303 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1304 | lemma filter_mset_eq_conv: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1305 | "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 | 1306 | proof | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1307 | 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 | 1308 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1309 | assume ?Q | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1310 | then obtain Q where M: "M = N + Q" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 1311 | by (auto simp add: mset_subset_eq_exists_conv) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1312 | then have MN: "M - N = Q" by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1313 | show ?P | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1314 | proof (rule multiset_eqI) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1315 | fix a | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1316 | 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 | 1317 | by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1318 | 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 | 1319 | proof (cases "a \<in># M") | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1320 | case True | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1321 | with * show ?thesis | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1322 | by (simp add: not_in_iff M) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1323 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1324 | case False then have "count M a = 0" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1325 | by (simp add: not_in_iff) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1326 | with M show ?thesis by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1327 | qed | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1328 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1329 | qed | 
| 59813 | 1330 | |
| 64077 | 1331 | lemma filter_filter_mset: "filter_mset P (filter_mset Q M) = {#x \<in># M. Q x \<and> P x#}"
 | 
| 1332 | by (auto simp: multiset_eq_iff) | |
| 1333 | ||
| 64418 | 1334 | lemma | 
| 1335 |   filter_mset_True[simp]: "{#y \<in># M. True#} = M" and
 | |
| 1336 |   filter_mset_False[simp]: "{#y \<in># M. False#} = {#}"
 | |
| 1337 | by (auto simp: multiset_eq_iff) | |
| 1338 | ||
| 75457 | 1339 | lemma filter_mset_cong0: | 
| 1340 | assumes "\<And>x. x \<in># M \<Longrightarrow> f x \<longleftrightarrow> g x" | |
| 1341 | shows "filter_mset f M = filter_mset g M" | |
| 1342 | proof (rule subset_mset.antisym; unfold subseteq_mset_def; rule allI) | |
| 1343 | fix x | |
| 1344 | show "count (filter_mset f M) x \<le> count (filter_mset g M) x" | |
| 1345 | using assms by (cases "x \<in># M") (simp_all add: not_in_iff) | |
| 1346 | next | |
| 1347 | fix x | |
| 1348 | show "count (filter_mset g M) x \<le> count (filter_mset f M) x" | |
| 1349 | using assms by (cases "x \<in># M") (simp_all add: not_in_iff) | |
| 1350 | qed | |
| 1351 | ||
| 1352 | lemma filter_mset_cong: | |
| 1353 | assumes "M = M'" and "\<And>x. x \<in># M' \<Longrightarrow> f x \<longleftrightarrow> g x" | |
| 1354 | shows "filter_mset f M = filter_mset g M'" | |
| 1355 | unfolding \<open>M = M'\<close> | |
| 1356 | using assms by (auto intro: filter_mset_cong0) | |
| 1357 | ||
| 79800 | 1358 | lemma filter_eq_replicate_mset: "{#y \<in># D. y = x#} = replicate_mset (count D x) x"
 | 
| 1359 | by (induct D) (simp add: multiset_eqI) | |
| 1360 | ||
| 59813 | 1361 | |
| 60500 | 1362 | subsubsection \<open>Size\<close> | 
| 10249 | 1363 | |
| 56656 | 1364 | definition wcount where "wcount f M = (\<lambda>x. count M x * Suc (f x))" | 
| 1365 | ||
| 1366 | lemma wcount_union: "wcount f (M + N) a = wcount f M a + wcount f N a" | |
| 1367 | by (auto simp: wcount_def add_mult_distrib) | |
| 1368 | ||
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1369 | lemma wcount_add_mset: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1370 | "wcount f (add_mset x M) a = (if x = a then Suc (f a) else 0) + wcount f M a" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1371 | unfolding add_mset_add_single[of _ M] wcount_union by (auto simp: wcount_def) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1372 | |
| 56656 | 1373 | definition size_multiset :: "('a \<Rightarrow> nat) \<Rightarrow> 'a multiset \<Rightarrow> nat" where
 | 
| 64267 | 1374 | "size_multiset f M = sum (wcount f M) (set_mset M)" | 
| 56656 | 1375 | |
| 1376 | lemmas size_multiset_eq = size_multiset_def[unfolded wcount_def] | |
| 1377 | ||
| 60606 | 1378 | instantiation multiset :: (type) size | 
| 1379 | begin | |
| 1380 | ||
| 56656 | 1381 | definition size_multiset where | 
| 1382 | 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 | 1383 | instance .. | 
| 60606 | 1384 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1385 | 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 | 1386 | |
| 56656 | 1387 | lemmas size_multiset_overloaded_eq = | 
| 1388 | size_multiset_overloaded_def[THEN fun_cong, unfolded size_multiset_eq, simplified] | |
| 1389 | ||
| 1390 | lemma size_multiset_empty [simp]: "size_multiset f {#} = 0"
 | |
| 80095 | 1391 | by (simp add: size_multiset_def) | 
| 56656 | 1392 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1393 | lemma size_empty [simp]: "size {#} = 0"
 | 
| 80095 | 1394 | by (simp add: size_multiset_overloaded_def) | 
| 56656 | 1395 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1396 | lemma size_multiset_single : "size_multiset f {#b#} = Suc (f b)"
 | 
| 80095 | 1397 | by (simp add: size_multiset_eq) | 
| 10249 | 1398 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1399 | lemma size_single: "size {#b#} = 1"
 | 
| 80095 | 1400 | by (simp add: size_multiset_overloaded_def size_multiset_single) | 
| 56656 | 1401 | |
| 64267 | 1402 | lemma sum_wcount_Int: | 
| 1403 | "finite A \<Longrightarrow> sum (wcount f N) (A \<inter> set_mset N) = sum (wcount f N) A" | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1404 | by (induct rule: finite_induct) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1405 | (simp_all add: Int_insert_left wcount_def count_eq_zero_iff) | 
| 56656 | 1406 | |
| 1407 | lemma size_multiset_union [simp]: | |
| 1408 | "size_multiset f (M + N::'a multiset) = size_multiset f M + size_multiset f N" | |
| 80095 | 1409 | apply (simp add: size_multiset_def sum_Un_nat sum.distrib sum_wcount_Int wcount_union) | 
| 1410 | by (metis add_implies_diff finite_set_mset inf.commute sum_wcount_Int) | |
| 10249 | 1411 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1412 | lemma size_multiset_add_mset [simp]: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1413 | "size_multiset f (add_mset a M) = Suc (f a) + size_multiset f M" | 
| 80095 | 1414 | by (metis add.commute add_mset_add_single size_multiset_single size_multiset_union) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1415 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1416 | lemma size_add_mset [simp]: "size (add_mset a A) = Suc (size A)" | 
| 80095 | 1417 | by (simp add: size_multiset_overloaded_def wcount_add_mset) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1418 | |
| 28708 
a1a436f09ec6
explicit check for pattern discipline before code translation
 haftmann parents: 
28562diff
changeset | 1419 | lemma size_union [simp]: "size (M + N::'a multiset) = size M + size N" | 
| 80095 | 1420 | by (auto simp add: size_multiset_overloaded_def) | 
| 56656 | 1421 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1422 | lemma size_multiset_eq_0_iff_empty [iff]: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1423 |   "size_multiset f M = 0 \<longleftrightarrow> M = {#}"
 | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1424 | by (auto simp add: size_multiset_eq count_eq_zero_iff) | 
| 10249 | 1425 | |
| 17161 | 1426 | lemma size_eq_0_iff_empty [iff]: "(size M = 0) = (M = {#})"
 | 
| 56656 | 1427 | by (auto simp add: size_multiset_overloaded_def) | 
| 26016 | 1428 | |
| 1429 | lemma nonempty_has_size: "(S \<noteq> {#}) = (0 < size S)"
 | |
| 80095 | 1430 | by (metis gr0I gr_implies_not0 size_empty size_eq_0_iff_empty) | 
| 10249 | 1431 | |
| 60607 | 1432 | lemma size_eq_Suc_imp_elem: "size M = Suc n \<Longrightarrow> \<exists>a. a \<in># M" | 
| 80095 | 1433 | using all_not_in_conv by fastforce | 
| 10249 | 1434 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1435 | 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 | 1436 | assumes "size M = Suc n" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1437 | shows "\<exists>a N. M = add_mset a N" | 
| 80095 | 1438 | by (metis assms insert_DiffM size_eq_Suc_imp_elem) | 
| 15869 | 1439 | |
| 60606 | 1440 | lemma size_mset_mono: | 
| 1441 | fixes A B :: "'a multiset" | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1442 | assumes "A \<subseteq># B" | 
| 60606 | 1443 | shows "size A \<le> size B" | 
| 59949 | 1444 | proof - | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 1445 | from assms[unfolded mset_subset_eq_exists_conv] | 
| 59949 | 1446 | obtain C where B: "B = A + C" by auto | 
| 60606 | 1447 | show ?thesis unfolding B by (induct C) auto | 
| 59949 | 1448 | qed | 
| 1449 | ||
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1450 | lemma size_filter_mset_lesseq[simp]: "size (filter_mset f M) \<le> size M" | 
| 80095 | 1451 | by (rule size_mset_mono[OF multiset_filter_subset]) | 
| 59949 | 1452 | |
| 1453 | lemma size_Diff_submset: | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1454 | "M \<subseteq># M' \<Longrightarrow> size (M' - M) = size M' - size(M::'a multiset)" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 1455 | by (metis add_diff_cancel_left' size_union mset_subset_eq_exists_conv) | 
| 26016 | 1456 | |
| 79800 | 1457 | lemma size_lt_imp_ex_count_lt: "size M < size N \<Longrightarrow> \<exists>x \<in># N. count M x < count N x" | 
| 1458 | by (metis count_eq_zero_iff leD not_le_imp_less not_less_zero size_mset_mono subseteq_mset_def) | |
| 1459 | ||
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1460 | |
| 60500 | 1461 | subsection \<open>Induction and case splits\<close> | 
| 10249 | 1462 | |
| 18258 | 1463 | theorem multiset_induct [case_names empty add, induct type: multiset]: | 
| 48009 | 1464 |   assumes empty: "P {#}"
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1465 | assumes add: "\<And>x M. P M \<Longrightarrow> P (add_mset x M)" | 
| 48009 | 1466 | shows "P M" | 
| 65545 | 1467 | proof (induct "size M" arbitrary: M) | 
| 48009 | 1468 | case 0 thus "P M" by (simp add: empty) | 
| 1469 | next | |
| 1470 | case (Suc k) | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1471 | obtain N x where "M = add_mset x N" | 
| 60500 | 1472 | using \<open>Suc k = size M\<close> [symmetric] | 
| 48009 | 1473 | using size_eq_Suc_imp_eq_union by fast | 
| 1474 | with Suc add show "P M" by simp | |
| 10249 | 1475 | qed | 
| 1476 | ||
| 65545 | 1477 | lemma multiset_induct_min[case_names empty add]: | 
| 1478 | fixes M :: "'a::linorder multiset" | |
| 1479 | assumes | |
| 1480 |     empty: "P {#}" and
 | |
| 1481 | add: "\<And>x M. P M \<Longrightarrow> (\<forall>y \<in># M. y \<ge> x) \<Longrightarrow> P (add_mset x M)" | |
| 1482 | shows "P M" | |
| 1483 | proof (induct "size M" arbitrary: M) | |
| 1484 | case (Suc k) | |
| 1485 | note ih = this(1) and Sk_eq_sz_M = this(2) | |
| 1486 | ||
| 66425 | 1487 | let ?y = "Min_mset M" | 
| 65545 | 1488 |   let ?N = "M - {#?y#}"
 | 
| 1489 | ||
| 1490 | have M: "M = add_mset ?y ?N" | |
| 1491 | by (metis Min_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero | |
| 1492 | set_mset_eq_empty_iff size_empty) | |
| 1493 | show ?case | |
| 1494 | by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, | |
| 1495 | meson Min_le finite_set_mset in_diffD) | |
| 1496 | qed (simp add: empty) | |
| 1497 | ||
| 1498 | lemma multiset_induct_max[case_names empty add]: | |
| 1499 | fixes M :: "'a::linorder multiset" | |
| 1500 | assumes | |
| 1501 |     empty: "P {#}" and
 | |
| 1502 | add: "\<And>x M. P M \<Longrightarrow> (\<forall>y \<in># M. y \<le> x) \<Longrightarrow> P (add_mset x M)" | |
| 1503 | shows "P M" | |
| 1504 | proof (induct "size M" arbitrary: M) | |
| 1505 | case (Suc k) | |
| 1506 | note ih = this(1) and Sk_eq_sz_M = this(2) | |
| 1507 | ||
| 66425 | 1508 | let ?y = "Max_mset M" | 
| 65545 | 1509 |   let ?N = "M - {#?y#}"
 | 
| 1510 | ||
| 1511 | have M: "M = add_mset ?y ?N" | |
| 1512 | by (metis Max_in Sk_eq_sz_M finite_set_mset insert_DiffM lessI not_less_zero | |
| 1513 | set_mset_eq_empty_iff size_empty) | |
| 1514 | show ?case | |
| 1515 | by (subst M, rule add, rule ih, metis M Sk_eq_sz_M nat.inject size_add_mset, | |
| 1516 | meson Max_ge finite_set_mset in_diffD) | |
| 1517 | qed (simp add: empty) | |
| 1518 | ||
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1519 | lemma multi_nonempty_split: "M \<noteq> {#} \<Longrightarrow> \<exists>A a. M = add_mset a A"
 | 
| 80095 | 1520 | by (induct M) auto | 
| 25610 | 1521 | |
| 55913 | 1522 | lemma multiset_cases [cases type]: | 
| 80095 | 1523 |   obtains (empty) "M = {#}" | (add) x N where "M = add_mset x N"
 | 
| 63092 | 1524 | by (induct M) simp_all | 
| 25610 | 1525 | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1526 | lemma multi_drop_mem_not_eq: "c \<in># B \<Longrightarrow> B - {#c#} \<noteq> B"
 | 
| 80095 | 1527 |   by (cases "B = {#}") (auto dest: multi_member_split)
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1528 | |
| 68992 | 1529 | lemma union_filter_mset_complement[simp]: | 
| 1530 | "\<forall>x. P x = (\<not> Q x) \<Longrightarrow> filter_mset P M + filter_mset Q M = M" | |
| 80095 | 1531 | by (subst multiset_eq_iff) auto | 
| 68992 | 1532 | |
| 66494 | 1533 | lemma multiset_partition: "M = {#x \<in># M. P x#} + {#x \<in># M. \<not> P x#}"
 | 
| 80095 | 1534 | by simp | 
| 66494 | 1535 | |
| 1536 | lemma mset_subset_size: "A \<subset># B \<Longrightarrow> size A < size B" | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1537 | proof (induct A arbitrary: B) | 
| 66494 | 1538 | case empty | 
| 1539 | then show ?case | |
| 1540 | using nonempty_has_size by auto | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1541 | next | 
| 66494 | 1542 | case (add x A) | 
| 1543 | have "add_mset x A \<subseteq># B" | |
| 1544 | by (meson add.prems subset_mset_def) | |
| 1545 | then show ?case | |
| 80095 | 1546 | using add.prems subset_mset.less_eqE by 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 | 1547 | 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 | 1548 | |
| 59949 | 1549 | lemma size_1_singleton_mset: "size M = 1 \<Longrightarrow> \<exists>a. M = {#a#}"
 | 
| 66494 | 1550 | by (cases M) auto | 
| 59949 | 1551 | |
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 1552 | lemma set_mset_subset_singletonD: | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 1553 |   assumes "set_mset A \<subseteq> {x}"
 | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 1554 | shows "A = replicate_mset (size A) x" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 1555 | using assms by (induction A) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 1556 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1557 | |
| 60500 | 1558 | subsubsection \<open>Strong induction and subset induction for multisets\<close> | 
| 1559 | ||
| 1560 | text \<open>Well-foundedness of strict subset relation\<close> | |
| 58098 | 1561 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 1562 | lemma wf_subset_mset_rel: "wf {(M, N :: 'a multiset). M \<subset># N}"
 | 
| 80322 | 1563 | using mset_subset_size wfp_def wfp_if_convertible_to_nat by blast | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1564 | |
| 80324 | 1565 | lemma wfp_subset_mset[simp]: "wfp (\<subset>#)" | 
| 76300 | 1566 | by (rule wf_subset_mset_rel[to_pred]) | 
| 1567 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1568 | lemma full_multiset_induct [case_names less]: | 
| 80095 | 1569 | assumes ih: "\<And>B. \<forall>(A::'a multiset). A \<subset># B \<longrightarrow> P A \<Longrightarrow> P B" | 
| 1570 | shows "P B" | |
| 1571 | apply (rule wf_subset_mset_rel [THEN wf_induct]) | |
| 1572 | apply (rule ih, auto) | |
| 1573 | 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 | 1574 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1575 | 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 | 1576 | assumes "F \<subseteq># A" | 
| 60606 | 1577 |     and empty: "P {#}"
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1578 | and insert: "\<And>a F. a \<in># A \<Longrightarrow> P F \<Longrightarrow> P (add_mset a F)" | 
| 60606 | 1579 | 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 | 1580 | proof - | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1581 | 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 | 1582 | 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 | 1583 | 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 | 1584 |     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 | 1585 | 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 | 1586 | fix x F | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1587 | assume P: "F \<subseteq># A \<Longrightarrow> P F" and i: "add_mset x F \<subseteq># A" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1588 | show "P (add_mset x F)" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1589 | proof (rule insert) | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 1590 | from i show "x \<in># A" by (auto dest: mset_subset_eq_insertD) | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 1591 | from i have "F \<subseteq># A" by (auto dest: mset_subset_eq_insertD) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1592 | 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 | 1593 | 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 | 1594 | 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 | 1595 | qed | 
| 26145 | 1596 | |
| 17161 | 1597 | |
| 75467 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1598 | subsection \<open>Least and greatest elements\<close> | 
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1599 | |
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1600 | context begin | 
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1601 | |
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1602 | qualified lemma | 
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1603 | assumes | 
| 77699 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 desharna parents: 
77688diff
changeset | 1604 |     "M \<noteq> {#}" and
 | 
| 76754 
b5f4ae037fe2
used transp_on in assumptions of lemmas Multiset.bex_(least|greatest)_element
 desharna parents: 
76749diff
changeset | 1605 | "transp_on (set_mset M) R" and | 
| 77699 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 desharna parents: 
77688diff
changeset | 1606 | "totalp_on (set_mset M) R" | 
| 75467 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1607 | shows | 
| 77699 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 desharna parents: 
77688diff
changeset | 1608 | bex_least_element: "(\<exists>l \<in># M. \<forall>x \<in># M. x \<noteq> l \<longrightarrow> R l x)" and | 
| 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 desharna parents: 
77688diff
changeset | 1609 | bex_greatest_element: "(\<exists>g \<in># M. \<forall>x \<in># M. x \<noteq> g \<longrightarrow> R x g)" | 
| 75467 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1610 | using assms | 
| 77699 
d5060a919b3f
reordered assumption and tuned proof of Multiset.bex_least_element and Multiset.bex_greatest_element
 desharna parents: 
77688diff
changeset | 1611 | by (auto intro: Finite_Set.bex_least_element Finite_Set.bex_greatest_element) | 
| 75467 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1612 | |
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1613 | end | 
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1614 | |
| 
9e34819a7ca1
added lemmas Multiset.bex_{least,greatest}_element
 desharna parents: 
75459diff
changeset | 1615 | |
| 60500 | 1616 | subsection \<open>The fold combinator\<close> | 
| 48023 | 1617 | |
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 1618 | definition fold_mset :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'a multiset \<Rightarrow> 'b"
 | 
| 48023 | 1619 | where | 
| 60495 | 1620 | "fold_mset f s M = Finite_Set.fold (\<lambda>x. f x ^^ count M x) s (set_mset M)" | 
| 48023 | 1621 | |
| 60606 | 1622 | 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 | 1623 | by (simp add: fold_mset_def) | 
| 48023 | 1624 | |
| 79800 | 1625 | lemma fold_mset_single [simp]: "fold_mset f s {#x#} = f x s"
 | 
| 1626 | by (simp add: fold_mset_def) | |
| 1627 | ||
| 48023 | 1628 | context comp_fun_commute | 
| 1629 | begin | |
| 1630 | ||
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1631 | lemma fold_mset_add_mset [simp]: "fold_mset f s (add_mset x M) = f x (fold_mset f s M)" | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1632 | proof - | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1633 | interpret mset: comp_fun_commute "\<lambda>y. f y ^^ count M y" | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1634 | by (fact comp_fun_commute_funpow) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1635 | interpret mset_union: comp_fun_commute "\<lambda>y. f y ^^ count (add_mset x M) y" | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1636 | by (fact comp_fun_commute_funpow) | 
| 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1637 | show ?thesis | 
| 60495 | 1638 | proof (cases "x \<in> set_mset M") | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1639 | case False | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1640 | then have *: "count (add_mset x M) x = 1" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1641 | by (simp add: not_in_iff) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1642 | from False have "Finite_Set.fold (\<lambda>y. f y ^^ count (add_mset x M) y) s (set_mset M) = | 
| 60495 | 1643 | Finite_Set.fold (\<lambda>y. f y ^^ count M y) s (set_mset M)" | 
| 73832 | 1644 | by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow) | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1645 | with False * show ?thesis | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1646 | by (simp add: fold_mset_def del: count_add_mset) | 
| 48023 | 1647 | next | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1648 | case True | 
| 63040 | 1649 |     define N where "N = set_mset M - {x}"
 | 
| 60495 | 1650 | from N_def True have *: "set_mset M = insert x N" "x \<notin> N" "finite N" by auto | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1651 | then have "Finite_Set.fold (\<lambda>y. f y ^^ count (add_mset x M) y) s N = | 
| 49822 
0cfc1651be25
simplified construction of fold combinator on multisets;
 haftmann parents: 
49717diff
changeset | 1652 | Finite_Set.fold (\<lambda>y. f y ^^ count M y) s N" | 
| 73832 | 1653 | by (auto intro!: Finite_Set.fold_cong comp_fun_commute_on_funpow) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1654 | with * show ?thesis by (simp add: fold_mset_def del: count_add_mset) simp | 
| 48023 | 1655 | qed | 
| 1656 | qed | |
| 1657 | ||
| 60606 | 1658 | lemma fold_mset_fun_left_comm: "f x (fold_mset f s M) = fold_mset f (f x s) M" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1659 | by (induct M) (simp_all add: fun_left_comm) | 
| 48023 | 1660 | |
| 60606 | 1661 | lemma fold_mset_union [simp]: "fold_mset f s (M + N) = fold_mset f (fold_mset f s M) N" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1662 | by (induct M) (simp_all add: fold_mset_fun_left_comm) | 
| 48023 | 1663 | |
| 1664 | lemma fold_mset_fusion: | |
| 1665 | assumes "comp_fun_commute g" | |
| 60606 | 1666 | and *: "\<And>x y. h (g x y) = f x (h y)" | 
| 1667 | shows "h (fold_mset g w A) = fold_mset f (h w) A" | |
| 48023 | 1668 | proof - | 
| 1669 | interpret comp_fun_commute g by (fact assms) | |
| 60606 | 1670 | from * show ?thesis by (induct A) auto | 
| 48023 | 1671 | qed | 
| 1672 | ||
| 1673 | end | |
| 1674 | ||
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1675 | lemma union_fold_mset_add_mset: "A + B = fold_mset add_mset A B" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1676 | proof - | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1677 | interpret comp_fun_commute add_mset | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1678 | by standard auto | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1679 | show ?thesis | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1680 | by (induction B) auto | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1681 | qed | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1682 | |
| 60500 | 1683 | text \<open> | 
| 48023 | 1684 | A note on code generation: When defining some function containing a | 
| 69593 | 1685 | subterm \<^term>\<open>fold_mset F\<close>, code generation is not automatic. When | 
| 61585 | 1686 | interpreting locale \<open>left_commutative\<close> with \<open>F\<close>, the | 
| 69593 | 1687 | would be code thms for \<^const>\<open>fold_mset\<close> become thms like | 
| 1688 |   \<^term>\<open>fold_mset F z {#} = z\<close> where \<open>F\<close> is not a pattern but
 | |
| 48023 | 1689 | contains defined symbols, i.e.\ is not a code thm. Hence a separate | 
| 61585 | 1690 | constant with its own code thms needs to be introduced for \<open>F\<close>. See the image operator below. | 
| 60500 | 1691 | \<close> | 
| 1692 | ||
| 1693 | ||
| 1694 | subsection \<open>Image\<close> | |
| 48023 | 1695 | |
| 1696 | definition image_mset :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset" where
 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1697 |   "image_mset f = fold_mset (add_mset \<circ> f) {#}"
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1698 | |
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1699 | lemma comp_fun_commute_mset_image: "comp_fun_commute (add_mset \<circ> f)" | 
| 66494 | 1700 | by unfold_locales (simp add: fun_eq_iff) | 
| 48023 | 1701 | |
| 1702 | lemma image_mset_empty [simp]: "image_mset f {#} = {#}"
 | |
| 49823 | 1703 | by (simp add: image_mset_def) | 
| 48023 | 1704 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1705 | lemma image_mset_single: "image_mset f {#x#} = {#f x#}"
 | 
| 66494 | 1706 | by (simp add: comp_fun_commute.fold_mset_add_mset comp_fun_commute_mset_image image_mset_def) | 
| 48023 | 1707 | |
| 60606 | 1708 | lemma image_mset_union [simp]: "image_mset f (M + N) = image_mset f M + image_mset f N" | 
| 49823 | 1709 | proof - | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1710 | interpret comp_fun_commute "add_mset \<circ> f" | 
| 49823 | 1711 | by (fact comp_fun_commute_mset_image) | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 1712 | show ?thesis by (induct N) (simp_all add: image_mset_def) | 
| 49823 | 1713 | qed | 
| 1714 | ||
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1715 | corollary image_mset_add_mset [simp]: | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1716 | "image_mset f (add_mset a M) = add_mset (f a) (image_mset f M)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1717 | unfolding image_mset_union add_mset_add_single[of a M] by (simp add: image_mset_single) | 
| 48023 | 1718 | |
| 60606 | 1719 | lemma set_image_mset [simp]: "set_mset (image_mset f M) = image f (set_mset M)" | 
| 49823 | 1720 | by (induct M) simp_all | 
| 48040 | 1721 | |
| 60606 | 1722 | lemma size_image_mset [simp]: "size (image_mset f M) = size M" | 
| 49823 | 1723 | by (induct M) simp_all | 
| 48023 | 1724 | |
| 60606 | 1725 | lemma image_mset_is_empty_iff [simp]: "image_mset f M = {#} \<longleftrightarrow> M = {#}"
 | 
| 49823 | 1726 | by (cases M) auto | 
| 48023 | 1727 | |
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1728 | lemma image_mset_If: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1729 | "image_mset (\<lambda>x. if P x then f x else g x) A = | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1730 | image_mset f (filter_mset P A) + image_mset g (filter_mset (\<lambda>x. \<not>P x) A)" | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 1731 | by (induction A) auto | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1732 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1733 | lemma image_mset_Diff: | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1734 | assumes "B \<subseteq># A" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1735 | shows "image_mset f (A - B) = image_mset f A - image_mset f B" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1736 | proof - | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1737 | have "image_mset f (A - B + B) = image_mset f (A - B) + image_mset f B" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1738 | by simp | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1739 | also from assms have "A - B + B = A" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1740 | by (simp add: subset_mset.diff_add) | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1741 | finally show ?thesis by simp | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1742 | qed | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1743 | |
| 80524 | 1744 | lemma minus_add_mset_if_not_in_lhs[simp]: "x \<notin># A \<Longrightarrow> A - add_mset x B = A - B" | 
| 1745 | by (metis diff_intersect_left_idem inter_add_right1) | |
| 1746 | ||
| 80525 | 1747 | lemma image_mset_diff_if_inj: | 
| 1748 | fixes f A B | |
| 1749 | assumes "inj f" | |
| 1750 | shows "image_mset f (A - B) = image_mset f A - image_mset f B" | |
| 1751 | proof (induction B) | |
| 1752 | case empty | |
| 1753 | show ?case | |
| 1754 | by simp | |
| 1755 | next | |
| 1756 | case (add x B) | |
| 1757 | show ?case | |
| 1758 | proof (cases "x \<in># A - B") | |
| 1759 | case True | |
| 1760 | ||
| 1761 | have "image_mset f (A - add_mset x B) = | |
| 1762 |         add_mset (f x) (image_mset f (A - add_mset x B)) - {#f x#}"
 | |
| 1763 | unfolding add_mset_remove_trivial .. | |
| 1764 | ||
| 1765 |     also have "\<dots> = image_mset f (add_mset x (A - add_mset x B)) - {#f x#}"
 | |
| 1766 | unfolding image_mset_add_mset .. | |
| 1767 | ||
| 1768 |     also have "\<dots> = image_mset f (add_mset x (A - B - {#x#})) - {#f x#}"
 | |
| 1769 | unfolding add_mset_add_single[symmetric] diff_diff_add_mset .. | |
| 1770 | ||
| 1771 |     also have "\<dots> = image_mset f (A - B) - {#f x#}"
 | |
| 1772 | unfolding insert_DiffM[OF \<open>x \<in># A - B\<close>] .. | |
| 1773 | ||
| 1774 |     also have "\<dots> = image_mset f A - image_mset f B - {#f x#}"
 | |
| 1775 | unfolding add.IH .. | |
| 1776 | ||
| 1777 | also have "\<dots> = image_mset f A - image_mset f (add_mset x B)" | |
| 1778 | unfolding diff_diff_add_mset add_mset_add_single[symmetric] image_mset_add_mset .. | |
| 1779 | ||
| 1780 | finally show ?thesis . | |
| 1781 | next | |
| 1782 | case False | |
| 1783 | ||
| 1784 | hence "image_mset f (A - add_mset x B) = image_mset f (A - B)" | |
| 1785 | using diff_single_trivial by fastforce | |
| 1786 | ||
| 1787 |     also have "\<dots> = image_mset f A - image_mset f B - {#f x#}"
 | |
| 1788 | proof - | |
| 1789 | have "f x \<notin> f ` set_mset (A - B)" | |
| 1790 | using False[folded inj_image_mem_iff[OF \<open>inj f\<close>]] . | |
| 1791 | ||
| 1792 | hence "f x \<notin># image_mset f (A - B)" | |
| 1793 | unfolding set_image_mset . | |
| 1794 | ||
| 1795 | thus ?thesis | |
| 1796 | unfolding add.IH[symmetric] | |
| 1797 | by (metis diff_single_trivial) | |
| 1798 | qed | |
| 1799 | ||
| 1800 | also have "\<dots> = image_mset f A - image_mset f (add_mset x B)" | |
| 1801 | by simp | |
| 1802 | ||
| 1803 | finally show ?thesis . | |
| 1804 | qed | |
| 1805 | qed | |
| 1806 | ||
| 73594 | 1807 | lemma count_image_mset: | 
| 1808 |   \<open>count (image_mset f A) x = (\<Sum>y\<in>f -` {x} \<inter> set_mset A. count A y)\<close>
 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1809 | proof (induction A) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1810 | case empty | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1811 | then show ?case by simp | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1812 | next | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1813 | case (add x A) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1814 | moreover have *: "(if x = y then Suc n else n) = n + (if x = y then 1 else 0)" for n y | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1815 | by simp | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1816 | ultimately show ?case | 
| 66494 | 1817 | by (auto simp: sum.distrib intro!: sum.mono_neutral_left) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1818 | qed | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 1819 | |
| 73594 | 1820 | lemma count_image_mset': | 
| 1821 | \<open>count (image_mset f X) y = (\<Sum>x | x \<in># X \<and> y = f x. count X x)\<close> | |
| 1822 | by (auto simp add: count_image_mset simp flip: singleton_conv2 simp add: Collect_conj_eq ac_simps) | |
| 1823 | ||
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 1824 | lemma image_mset_subseteq_mono: "A \<subseteq># B \<Longrightarrow> image_mset f A \<subseteq># image_mset f B" | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 1825 | by (metis image_mset_union subset_mset.le_iff_add) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 1826 | |
| 65048 | 1827 | lemma image_mset_subset_mono: "M \<subset># N \<Longrightarrow> image_mset f M \<subset># image_mset f N" | 
| 1828 | by (metis (no_types) Diff_eq_empty_iff_mset image_mset_Diff image_mset_is_empty_iff | |
| 1829 | image_mset_subseteq_mono subset_mset.less_le_not_le) | |
| 1830 | ||
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1831 | syntax (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1832 |   "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ :# _#})")
 | 
| 48023 | 1833 | syntax | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1834 |   "_comprehension_mset" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> 'a multiset"  ("({#_/. _ \<in># _#})")
 | 
| 80768 | 1835 | syntax_consts | 
| 1836 | "_comprehension_mset" \<rightleftharpoons> image_mset | |
| 59813 | 1837 | translations | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1838 |   "{#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 | 1839 | |
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1840 | syntax (ASCII) | 
| 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1841 |   "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ :# _./ _#})")
 | 
| 48023 | 1842 | syntax | 
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 1843 |   "_comprehension_mset'" :: "'a \<Rightarrow> 'b \<Rightarrow> 'b multiset \<Rightarrow> bool \<Rightarrow> 'a multiset"  ("({#_/ | _ \<in># _./ _#})")
 | 
| 80768 | 1844 | syntax_consts | 
| 1845 | "_comprehension_mset'" \<rightleftharpoons> image_mset | |
| 59813 | 1846 | translations | 
| 60606 | 1847 |   "{#e | x\<in>#M. P#}" \<rightharpoonup> "{#e. x \<in># {# x\<in>#M. P#}#}"
 | 
| 59813 | 1848 | |
| 60500 | 1849 | text \<open> | 
| 69593 | 1850 |   This allows to write not just filters like \<^term>\<open>{#x\<in>#M. x<c#}\<close>
 | 
| 1851 |   but also images like \<^term>\<open>{#x+x. x\<in>#M #}\<close> and @{term [source]
 | |
| 60607 | 1852 |   "{#x+x|x\<in>#M. x<c#}"}, where the latter is currently displayed as
 | 
| 69593 | 1853 |   \<^term>\<open>{#x+x|x\<in>#M. x<c#}\<close>.
 | 
| 60500 | 1854 | \<close> | 
| 48023 | 1855 | |
| 60495 | 1856 | lemma in_image_mset: "y \<in># {#f x. x \<in># M#} \<longleftrightarrow> y \<in> f ` set_mset M"
 | 
| 66494 | 1857 | by simp | 
| 59813 | 1858 | |
| 55467 
a5c9002bc54d
renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
 blanchet parents: 
55417diff
changeset | 1859 | functor image_mset: image_mset | 
| 48023 | 1860 | proof - | 
| 1861 | fix f g show "image_mset f \<circ> image_mset g = image_mset (f \<circ> g)" | |
| 1862 | proof | |
| 1863 | fix A | |
| 1864 | show "(image_mset f \<circ> image_mset g) A = image_mset (f \<circ> g) A" | |
| 1865 | by (induct A) simp_all | |
| 1866 | qed | |
| 1867 | show "image_mset id = id" | |
| 1868 | proof | |
| 1869 | fix A | |
| 1870 | show "image_mset id A = id A" | |
| 1871 | by (induct A) simp_all | |
| 1872 | qed | |
| 1873 | qed | |
| 1874 | ||
| 59813 | 1875 | declare | 
| 1876 | image_mset.id [simp] | |
| 1877 | image_mset.identity [simp] | |
| 1878 | ||
| 1879 | lemma image_mset_id[simp]: "image_mset id x = x" | |
| 1880 | unfolding id_def by auto | |
| 1881 | ||
| 1882 | 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#}"
 | |
| 1883 | by (induct M) auto | |
| 1884 | ||
| 1885 | lemma image_mset_cong_pair: | |
| 1886 |   "(\<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#}"
 | |
| 1887 | by (metis image_mset_cong split_cong) | |
| 49717 | 1888 | |
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64587diff
changeset | 1889 | lemma image_mset_const_eq: | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64587diff
changeset | 1890 |   "{#c. a \<in># M#} = replicate_mset (size M) c"
 | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64587diff
changeset | 1891 | by (induct M) simp_all | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64587diff
changeset | 1892 | |
| 75459 | 1893 | lemma image_mset_filter_mset_swap: | 
| 1894 | "image_mset f (filter_mset (\<lambda>x. P (f x)) M) = filter_mset P (image_mset f M)" | |
| 1895 | by (induction M rule: multiset_induct) simp_all | |
| 1896 | ||
| 75560 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1897 | lemma image_mset_eq_plusD: | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1898 | "image_mset f A = B + C \<Longrightarrow> \<exists>B' C'. A = B' + C' \<and> B = image_mset f B' \<and> C = image_mset f C'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1899 | proof (induction A arbitrary: B C) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1900 | case empty | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1901 | thus ?case by simp | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1902 | next | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1903 | case (add x A) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1904 | show ?case | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1905 | proof (cases "f x \<in># B") | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1906 | case True | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1907 |     with add.prems have "image_mset f A = (B - {#f x#}) + C"
 | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1908 | by (metis add_mset_remove_trivial image_mset_add_mset mset_subset_eq_single | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1909 | subset_mset.add_diff_assoc2) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1910 | thus ?thesis | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1911 | using add.IH add.prems by force | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1912 | next | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1913 | case False | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1914 |     with add.prems have "image_mset f A = B + (C - {#f x#})"
 | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1915 | by (metis diff_single_eq_union diff_union_single_conv image_mset_add_mset union_iff | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1916 | union_single_eq_member) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1917 | then show ?thesis | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1918 | using add.IH add.prems by force | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1919 | qed | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1920 | qed | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1921 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1922 | lemma image_mset_eq_image_mset_plusD: | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1923 | assumes "image_mset f A = image_mset f B + C" and inj_f: "inj_on f (set_mset A \<union> set_mset B)" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1924 | shows "\<exists>C'. A = B + C' \<and> C = image_mset f C'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1925 | using assms | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1926 | proof (induction A arbitrary: B C) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1927 | case empty | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1928 | thus ?case by simp | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1929 | next | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1930 | case (add x A) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1931 | show ?case | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1932 | proof (cases "x \<in># B") | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1933 | case True | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1934 |     with add.prems have "image_mset f A = image_mset f (B - {#x#}) + C"
 | 
| 80095 | 1935 | by (smt (verit) add_mset_add_mset_same_iff image_mset_add_mset insert_DiffM union_mset_add_mset_left) | 
| 75560 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1936 |     with add.IH have "\<exists>M3'. A = B - {#x#} + M3' \<and> image_mset f M3' = C"
 | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1937 | by (smt (verit, del_insts) True Un_insert_left Un_insert_right add.prems(2) inj_on_insert | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1938 | insert_DiffM set_mset_add_mset_insert) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1939 | with True show ?thesis | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1940 | by auto | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1941 | next | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1942 | case False | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1943 | with add.prems(2) have "f x \<notin># image_mset f B" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1944 | by auto | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1945 |     with add.prems(1) have "image_mset f A = image_mset f B + (C - {#f x#})"
 | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1946 | by (metis (no_types, lifting) diff_union_single_conv image_eqI image_mset_Diff | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1947 | image_mset_single mset_subset_eq_single set_image_mset union_iff union_single_eq_diff | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1948 | union_single_eq_member) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1949 |     with add.prems(2) add.IH have "\<exists>M3'. A = B + M3' \<and> C - {#f x#} = image_mset f M3'"
 | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1950 | by auto | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1951 | then show ?thesis | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1952 | by (metis add.prems(1) add_diff_cancel_left' image_mset_Diff mset_subset_eq_add_left | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1953 | union_mset_add_mset_right) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1954 | qed | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1955 | qed | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1956 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1957 | lemma image_mset_eq_plus_image_msetD: | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1958 | "image_mset f A = B + image_mset f C \<Longrightarrow> inj_on f (set_mset A \<union> set_mset C) \<Longrightarrow> | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1959 | \<exists>B'. A = B' + C \<and> B = image_mset f B'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1960 | unfolding add.commute[of B] add.commute[of _ C] | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1961 | by (rule image_mset_eq_image_mset_plusD; assumption) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1962 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 1963 | |
| 60500 | 1964 | 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 | 1965 | |
| 60515 | 1966 | primrec mset :: "'a list \<Rightarrow> 'a multiset" where | 
| 1967 |   "mset [] = {#}" |
 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 1968 | "mset (a # x) = add_mset a (mset x)" | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1969 | |
| 37107 | 1970 | lemma in_multiset_in_set: | 
| 60515 | 1971 | "x \<in># mset xs \<longleftrightarrow> x \<in> set xs" | 
| 37107 | 1972 | by (induct xs) simp_all | 
| 1973 | ||
| 60515 | 1974 | lemma count_mset: | 
| 1975 | "count (mset xs) x = length (filter (\<lambda>y. x = y) xs)" | |
| 37107 | 1976 | by (induct xs) simp_all | 
| 1977 | ||
| 60515 | 1978 | lemma mset_zero_iff[simp]: "(mset x = {#}) = (x = [])"
 | 
| 59813 | 1979 | 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 | 1980 | |
| 60515 | 1981 | 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 | 1982 | 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 | 1983 | |
| 66276 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 1984 | lemma count_mset_gt_0: "x \<in> set xs \<Longrightarrow> count (mset xs) x > 0" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 1985 | by (induction xs) auto | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 1986 | |
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 1987 | lemma count_mset_0_iff [simp]: "count (mset xs) x = 0 \<longleftrightarrow> x \<notin> set xs" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 1988 | by (induction xs) auto | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 1989 | |
| 64077 | 1990 | lemma mset_single_iff[iff]: "mset xs = {#x#} \<longleftrightarrow> xs = [x]"
 | 
| 1991 | by (cases xs) auto | |
| 1992 | ||
| 1993 | lemma mset_single_iff_right[iff]: "{#x#} = mset xs \<longleftrightarrow> xs = [x]"
 | |
| 1994 | by (cases xs) auto | |
| 1995 | ||
| 64076 | 1996 | lemma set_mset_mset[simp]: "set_mset (mset xs) = set xs" | 
| 1997 | by (induct xs) auto | |
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 1998 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 1999 | 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 | 2000 | 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 | 2001 | |
| 60515 | 2002 | lemma size_mset [simp]: "size (mset xs) = length xs" | 
| 48012 | 2003 | by (induct xs) simp_all | 
| 2004 | ||
| 60606 | 2005 | lemma mset_append [simp]: "mset (xs @ ys) = mset xs + mset ys" | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2006 | by (induct xs arbitrary: ys) auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2007 | |
| 68988 | 2008 | lemma mset_filter[simp]: "mset (filter P xs) = {#x \<in># mset xs. P x #}"
 | 
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 2009 | by (induct xs) simp_all | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 2010 | |
| 60515 | 2011 | lemma mset_rev [simp]: | 
| 2012 | "mset (rev xs) = mset xs" | |
| 40950 | 2013 | by (induct xs) simp_all | 
| 2014 | ||
| 60515 | 2015 | lemma surj_mset: "surj mset" | 
| 76359 | 2016 | unfolding surj_def | 
| 2017 | proof (rule allI) | |
| 2018 | fix M | |
| 2019 | show "\<exists>xs. M = mset xs" | |
| 2020 | by (induction M) (auto intro: exI[of _ "_ # _"]) | |
| 2021 | 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 | 2022 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 2023 | lemma distinct_count_atmost_1: | 
| 60606 | 2024 | "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 | 2025 | proof (induct x) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2026 | case Nil then show ?case by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2027 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2028 | 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 | 2029 | proof | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2030 | assume ?lhs then show ?rhs using Cons by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2031 | next | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2032 | assume ?rhs then have "x \<notin> set xs" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2033 | by (simp split: if_splits) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2034 | 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 | 2035 | (if a \<in> set xs then 1 else 0))" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2036 | 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 | 2037 | ultimately show ?lhs using Cons by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2038 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2039 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2040 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2041 | lemma mset_eq_setD: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2042 | assumes "mset xs = mset ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2043 | shows "set xs = set ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2044 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2045 | 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 | 2046 | by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2047 | then show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2048 | 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 | 2049 | |
| 60515 | 2050 | lemma set_eq_iff_mset_eq_distinct: | 
| 73301 | 2051 | \<open>distinct x \<Longrightarrow> distinct y \<Longrightarrow> set x = set y \<longleftrightarrow> mset x = mset y\<close> | 
| 2052 | 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 | 2053 | |
| 60515 | 2054 | lemma set_eq_iff_mset_remdups_eq: | 
| 73301 | 2055 | \<open>set x = set y \<longleftrightarrow> mset (remdups x) = mset (remdups y)\<close> | 
| 80095 | 2056 | using set_eq_iff_mset_eq_distinct by 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 | 2057 | |
| 73301 | 2058 | lemma mset_eq_imp_distinct_iff: | 
| 2059 | \<open>distinct xs \<longleftrightarrow> distinct ys\<close> if \<open>mset xs = mset ys\<close> | |
| 2060 | using that by (auto simp add: distinct_count_atmost_1 dest: mset_eq_setD) | |
| 2061 | ||
| 60607 | 2062 | lemma nth_mem_mset: "i < length ls \<Longrightarrow> (ls ! i) \<in># mset ls" | 
| 60678 | 2063 | proof (induct ls arbitrary: i) | 
| 2064 | case Nil | |
| 2065 | then show ?case by simp | |
| 2066 | next | |
| 2067 | case Cons | |
| 2068 | then show ?case by (cases i) auto | |
| 2069 | 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 | 2070 | |
| 60606 | 2071 | lemma mset_remove1[simp]: "mset (remove1 a xs) = mset xs - {#a#}"
 | 
| 60678 | 2072 | 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 | 2073 | |
| 60515 | 2074 | lemma mset_eq_length: | 
| 2075 | assumes "mset xs = mset ys" | |
| 37107 | 2076 | shows "length xs = length ys" | 
| 60515 | 2077 | using assms by (metis size_mset) | 
| 2078 | ||
| 2079 | lemma mset_eq_length_filter: | |
| 2080 | 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 | 2081 | shows "length (filter (\<lambda>x. z = x) xs) = length (filter (\<lambda>y. z = y) ys)" | 
| 60515 | 2082 | 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 | 2083 | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 2084 | lemma fold_multiset_equiv: | 
| 73706 | 2085 | \<open>List.fold f xs = List.fold f ys\<close> | 
| 2086 | if f: \<open>\<And>x y. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> f x \<circ> f y = f y \<circ> f x\<close> | |
| 2087 | and \<open>mset xs = mset ys\<close> | |
| 2088 | using f \<open>mset xs = mset ys\<close> [symmetric] proof (induction xs arbitrary: ys) | |
| 60678 | 2089 | case Nil | 
| 2090 | then show ?case by simp | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 2091 | next | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 2092 | case (Cons x xs) | 
| 73706 | 2093 | then have *: \<open>set ys = set (x # xs)\<close> | 
| 60678 | 2094 | by (blast dest: mset_eq_setD) | 
| 73706 | 2095 | have \<open>\<And>x y. x \<in> set ys \<Longrightarrow> y \<in> set ys \<Longrightarrow> f x \<circ> f y = f y \<circ> f x\<close> | 
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 2096 | by (rule Cons.prems(1)) (simp_all add: *) | 
| 73706 | 2097 | moreover from * have \<open>x \<in> set ys\<close> | 
| 2098 | by simp | |
| 2099 | ultimately have \<open>List.fold f ys = List.fold f (remove1 x ys) \<circ> f x\<close> | |
| 2100 | by (fact fold_remove1_split) | |
| 2101 | moreover from Cons.prems have \<open>List.fold f xs = List.fold f (remove1 x ys)\<close> | |
| 2102 | by (auto intro: Cons.IH) | |
| 2103 | ultimately show ?case | |
| 60678 | 2104 | by simp | 
| 73706 | 2105 | qed | 
| 2106 | ||
| 2107 | lemma fold_permuted_eq: | |
| 2108 | \<open>List.fold (\<odot>) xs z = List.fold (\<odot>) ys z\<close> | |
| 2109 | if \<open>mset xs = mset ys\<close> | |
| 2110 | and \<open>P z\<close> and P: \<open>\<And>x z. x \<in> set xs \<Longrightarrow> P z \<Longrightarrow> P (x \<odot> z)\<close> | |
| 2111 | and f: \<open>\<And>x y z. x \<in> set xs \<Longrightarrow> y \<in> set xs \<Longrightarrow> P z \<Longrightarrow> x \<odot> (y \<odot> z) = y \<odot> (x \<odot> z)\<close> | |
| 2112 | for f (infixl \<open>\<odot>\<close> 70) | |
| 2113 | using \<open>P z\<close> P f \<open>mset xs = mset ys\<close> [symmetric] proof (induction xs arbitrary: ys z) | |
| 2114 | case Nil | |
| 2115 | then show ?case by simp | |
| 2116 | next | |
| 2117 | case (Cons x xs) | |
| 2118 | then have *: \<open>set ys = set (x # xs)\<close> | |
| 2119 | by (blast dest: mset_eq_setD) | |
| 2120 | have \<open>P z\<close> | |
| 2121 | by (fact Cons.prems(1)) | |
| 2122 | moreover have \<open>\<And>x z. x \<in> set ys \<Longrightarrow> P z \<Longrightarrow> P (x \<odot> z)\<close> | |
| 2123 | by (rule Cons.prems(2)) (simp_all add: *) | |
| 2124 | moreover have \<open>\<And>x y z. x \<in> set ys \<Longrightarrow> y \<in> set ys \<Longrightarrow> P z \<Longrightarrow> x \<odot> (y \<odot> z) = y \<odot> (x \<odot> z)\<close> | |
| 2125 | by (rule Cons.prems(3)) (simp_all add: *) | |
| 2126 | moreover from * have \<open>x \<in> set ys\<close> | |
| 2127 | by simp | |
| 2128 | ultimately have \<open>fold (\<odot>) ys z = fold (\<odot>) (remove1 x ys) (x \<odot> z)\<close> | |
| 2129 | by (induction ys arbitrary: z) auto | |
| 2130 | moreover from Cons.prems have \<open>fold (\<odot>) xs (x \<odot> z) = fold (\<odot>) (remove1 x ys) (x \<odot> z)\<close> | |
| 2131 | by (auto intro: Cons.IH) | |
| 2132 | ultimately show ?case | |
| 2133 | by simp | |
| 45989 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 2134 | qed | 
| 
b39256df5f8a
moved theorem requiring multisets from More_List to Multiset
 haftmann parents: 
45866diff
changeset | 2135 | |
| 69107 | 2136 | lemma mset_shuffles: "zs \<in> shuffles xs ys \<Longrightarrow> mset zs = mset xs + mset ys" | 
| 2137 | by (induction xs ys arbitrary: zs rule: shuffles.induct) auto | |
| 65350 
b149abe619f7
added shuffle product to HOL/List
 eberlm <eberlm@in.tum.de> parents: 
65048diff
changeset | 2138 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2139 | lemma mset_insort [simp]: "mset (insort x xs) = add_mset x (mset xs)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2140 | by (induct xs) simp_all | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2141 | |
| 63524 
4ec755485732
adding mset_map to the simp rules
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63410diff
changeset | 2142 | lemma mset_map[simp]: "mset (map f xs) = image_mset f (mset xs)" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2143 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2144 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2145 | global_interpretation mset_set: folding add_mset "{#}"
 | 
| 73832 | 2146 |   defines mset_set = "folding_on.F add_mset {#}"
 | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2147 | by standard (simp add: fun_eq_iff) | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2148 | |
| 66276 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2149 | lemma sum_multiset_singleton [simp]: "sum (\<lambda>n. {#n#}) A = mset_set A"
 | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2150 | by (induction A rule: infinite_finite_induct) auto | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2151 | |
| 60513 | 2152 | lemma count_mset_set [simp]: | 
| 2153 | "finite A \<Longrightarrow> x \<in> A \<Longrightarrow> count (mset_set A) x = 1" (is "PROP ?P") | |
| 2154 | "\<not> finite A \<Longrightarrow> count (mset_set A) x = 0" (is "PROP ?Q") | |
| 2155 | "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 | 2156 | proof - | 
| 60606 | 2157 | have *: "count (mset_set A) x = 0" if "x \<notin> A" for A | 
| 2158 | proof (cases "finite A") | |
| 2159 | case False then show ?thesis by simp | |
| 2160 | next | |
| 2161 | case True from True \<open>x \<notin> A\<close> show ?thesis by (induct A) auto | |
| 2162 | qed | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 2163 | 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 | 2164 | by (auto elim!: Set.set_insert) | 
| 69593 | 2165 | qed \<comment> \<open>TODO: maybe define \<^const>\<open>mset_set\<close> also in terms of \<^const>\<open>Abs_multiset\<close>\<close> | 
| 60513 | 2166 | |
| 2167 | lemma elem_mset_set[simp, intro]: "finite A \<Longrightarrow> x \<in># mset_set A \<longleftrightarrow> x \<in> A" | |
| 59813 | 2168 | by (induct A rule: finite_induct) simp_all | 
| 2169 | ||
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2170 | lemma mset_set_Union: | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2171 |   "finite A \<Longrightarrow> finite B \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> mset_set (A \<union> B) = mset_set A + mset_set B"
 | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2172 | by (induction A rule: finite_induct) auto | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2173 | |
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2174 | lemma filter_mset_mset_set [simp]: | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2175 |   "finite A \<Longrightarrow> filter_mset P (mset_set A) = mset_set {x\<in>A. P x}"
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2176 | proof (induction A rule: finite_induct) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2177 | case (insert x A) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2178 | from insert.hyps have "filter_mset P (mset_set (insert x A)) = | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2179 |       filter_mset P (mset_set A) + mset_set (if P x then {x} else {})"
 | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2180 | by simp | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2181 |   also have "filter_mset P (mset_set A) = mset_set {x\<in>A. P x}"
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2182 | by (rule insert.IH) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2183 | also from insert.hyps | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2184 |     have "\<dots> + mset_set (if P x then {x} else {}) =
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2185 |             mset_set ({x \<in> A. P x} \<union> (if P x then {x} else {}))" (is "_ = mset_set ?A")
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2186 | by (intro mset_set_Union [symmetric]) simp_all | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2187 |   also from insert.hyps have "?A = {y\<in>insert x A. P y}" by auto
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2188 | finally show ?case . | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2189 | qed simp_all | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2190 | |
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2191 | lemma mset_set_Diff: | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2192 | assumes "finite A" "B \<subseteq> A" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2193 | shows "mset_set (A - B) = mset_set A - mset_set B" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2194 | proof - | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2195 | from assms have "mset_set ((A - B) \<union> B) = mset_set (A - B) + mset_set B" | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2196 | by (intro mset_set_Union) (auto dest: finite_subset) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2197 | also from assms have "A - B \<union> B = A" by blast | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2198 | finally show ?thesis by simp | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2199 | qed | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2200 | |
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2201 | lemma mset_set_set: "distinct xs \<Longrightarrow> mset_set (set xs) = mset xs" | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2202 | by (induction xs) simp_all | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2203 | |
| 66276 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2204 | lemma count_mset_set': "count (mset_set A) x = (if finite A \<and> x \<in> A then 1 else 0)" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2205 | by auto | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2206 | |
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2207 | lemma subset_imp_msubset_mset_set: | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2208 | assumes "A \<subseteq> B" "finite B" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2209 | shows "mset_set A \<subseteq># mset_set B" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2210 | proof (rule mset_subset_eqI) | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2211 | fix x :: 'a | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2212 | from assms have "finite A" by (rule finite_subset) | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2213 | with assms show "count (mset_set A) x \<le> count (mset_set B) x" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2214 | by (cases "x \<in> A"; cases "x \<in> B") auto | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2215 | qed | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2216 | |
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2217 | lemma mset_set_set_mset_msubset: "mset_set (set_mset A) \<subseteq># A" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2218 | proof (rule mset_subset_eqI) | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2219 | fix x show "count (mset_set (set_mset A)) x \<le> count A x" | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2220 | by (cases "x \<in># A") simp_all | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2221 | qed | 
| 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2222 | |
| 73466 | 2223 | lemma mset_set_upto_eq_mset_upto: | 
| 2224 |   \<open>mset_set {..<n} = mset [0..<n]\<close>
 | |
| 2225 | by (induction n) (auto simp: ac_simps lessThan_Suc) | |
| 2226 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2227 | context linorder | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2228 | begin | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2229 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2230 | definition sorted_list_of_multiset :: "'a multiset \<Rightarrow> 'a list" | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2231 | where | 
| 59998 
c54d36be22ef
renamed Multiset.fold -> fold_mset, Multiset.filter -> filter_mset
 nipkow parents: 
59986diff
changeset | 2232 | "sorted_list_of_multiset M = fold_mset insort [] M" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2233 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2234 | lemma sorted_list_of_multiset_empty [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2235 |   "sorted_list_of_multiset {#} = []"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2236 | by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2237 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2238 | lemma sorted_list_of_multiset_singleton [simp]: | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2239 |   "sorted_list_of_multiset {#x#} = [x]"
 | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2240 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2241 | interpret comp_fun_commute insort by (fact comp_fun_commute_insort) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2242 | show ?thesis by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2243 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2244 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2245 | lemma sorted_list_of_multiset_insert [simp]: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2246 | "sorted_list_of_multiset (add_mset x M) = List.insort x (sorted_list_of_multiset M)" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2247 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2248 | interpret comp_fun_commute insort by (fact comp_fun_commute_insort) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2249 | show ?thesis by (simp add: sorted_list_of_multiset_def) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2250 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2251 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2252 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2253 | |
| 66494 | 2254 | lemma mset_sorted_list_of_multiset[simp]: "mset (sorted_list_of_multiset M) = M" | 
| 2255 | by (induct M) simp_all | |
| 2256 | ||
| 2257 | lemma sorted_list_of_multiset_mset[simp]: "sorted_list_of_multiset (mset xs) = sort xs" | |
| 2258 | by (induct xs) simp_all | |
| 2259 | ||
| 2260 | lemma finite_set_mset_mset_set[simp]: "finite A \<Longrightarrow> set_mset (mset_set A) = A" | |
| 2261 | by auto | |
| 60513 | 2262 | |
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2263 | lemma mset_set_empty_iff: "mset_set A = {#} \<longleftrightarrow> A = {} \<or> infinite A"
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2264 | using finite_set_mset_mset_set by fastforce | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2265 | |
| 66494 | 2266 | lemma infinite_set_mset_mset_set: "\<not> finite A \<Longrightarrow> set_mset (mset_set A) = {}"
 | 
| 2267 | by simp | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2268 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2269 | lemma set_sorted_list_of_multiset [simp]: | 
| 60495 | 2270 | "set (sorted_list_of_multiset M) = set_mset M" | 
| 80095 | 2271 | by (induct M) (simp_all add: set_insort_key) | 
| 60513 | 2272 | |
| 2273 | lemma sorted_list_of_mset_set [simp]: | |
| 2274 | "sorted_list_of_multiset (mset_set A) = sorted_list_of_set A" | |
| 80095 | 2275 | by (cases "finite A") (induct A rule: finite_induct, simp_all) | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2276 | |
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2277 | lemma mset_upt [simp]: "mset [m..<n] = mset_set {m..<n}"
 | 
| 80095 | 2278 | by (metis distinct_upt mset_set_set set_upt) | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2279 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2280 | lemma image_mset_map_of: | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2281 |   "distinct (map fst xs) \<Longrightarrow> {#the (map_of xs i). i \<in># mset (map fst xs)#} = mset (map snd xs)"
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2282 | proof (induction xs) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2283 | case (Cons x xs) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2284 |   have "{#the (map_of (x # xs) i). i \<in># mset (map fst (x # xs))#} =
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2285 |           add_mset (snd x) {#the (if i = fst x then Some (snd x) else map_of xs i).
 | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2286 | i \<in># mset (map fst xs)#}" (is "_ = add_mset _ ?A") by simp | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2287 |   also from Cons.prems have "?A = {#the (map_of xs i). i :# mset (map fst xs)#}"
 | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2288 | by (cases x, intro image_mset_cong) (auto simp: in_multiset_in_set) | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2289 | also from Cons.prems have "\<dots> = mset (map snd xs)" by (intro Cons.IH) simp_all | 
| 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2290 | finally show ?case by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2291 | qed simp_all | 
| 63099 
af0e964aad7b
Moved material from AFP/Randomised_Social_Choice to distribution
 eberlm parents: 
63092diff
changeset | 2292 | |
| 66494 | 2293 | lemma msubset_mset_set_iff[simp]: | 
| 66276 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2294 | assumes "finite A" "finite B" | 
| 66494 | 2295 | shows "mset_set A \<subseteq># mset_set B \<longleftrightarrow> A \<subseteq> B" | 
| 2296 | using assms set_mset_mono subset_imp_msubset_mset_set by fastforce | |
| 2297 | ||
| 2298 | lemma mset_set_eq_iff[simp]: | |
| 2299 | assumes "finite A" "finite B" | |
| 2300 | shows "mset_set A = mset_set B \<longleftrightarrow> A = B" | |
| 2301 | using assms by (fastforce dest: finite_set_mset_mset_set) | |
| 66276 
acc3b7dd0b21
More material on powers for HOL-Computational_Algebra/HOL-Number_Theory
 eberlm <eberlm@in.tum.de> parents: 
65547diff
changeset | 2302 | |
| 69895 
6b03a8cf092d
more formal contributors (with the help of the history);
 wenzelm parents: 
69605diff
changeset | 2303 | lemma image_mset_mset_set: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close> | 
| 63921 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2304 | assumes "inj_on f A" | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2305 | shows "image_mset f (mset_set A) = mset_set (f ` A)" | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2306 | proof cases | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2307 | assume "finite A" | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2308 | from this \<open>inj_on f A\<close> show ?thesis | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2309 | by (induct A) auto | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2310 | next | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2311 | assume "infinite A" | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2312 | from this \<open>inj_on f A\<close> have "infinite (f ` A)" | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2313 | using finite_imageD by blast | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2314 | from \<open>infinite A\<close> \<open>infinite (f ` A)\<close> show ?thesis by simp | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2315 | qed | 
| 
0a5184877cb7
Additions to permutations (contributed by Lukas Bulwahn)
 eberlm <eberlm@in.tum.de> parents: 
63908diff
changeset | 2316 | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2317 | |
| 79800 | 2318 | subsection \<open>More properties of the replicate, repeat, and image operations\<close> | 
| 60804 | 2319 | |
| 2320 | 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 | 2321 | unfolding replicate_mset_def by (induct n) auto | 
| 60804 | 2322 | |
| 2323 | lemma set_mset_replicate_mset_subset[simp]: "set_mset (replicate_mset n x) = (if n = 0 then {} else {x})"
 | |
| 80095 | 2324 | by auto | 
| 60804 | 2325 | |
| 2326 | lemma size_replicate_mset[simp]: "size (replicate_mset n M) = n" | |
| 2327 | by (induct n, simp_all) | |
| 2328 | ||
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2329 | lemma size_repeat_mset [simp]: "size (repeat_mset n A) = n * size A" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2330 | by (induction n) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2331 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2332 | lemma size_multiset_sum [simp]: "size (\<Sum>x\<in>A. f x :: 'a multiset) = (\<Sum>x\<in>A. size (f x))" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2333 | by (induction A rule: infinite_finite_induct) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2334 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2335 | lemma size_multiset_sum_list [simp]: "size (\<Sum>X\<leftarrow>Xs. X :: 'a multiset) = (\<Sum>X\<leftarrow>Xs. size X)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2336 | by (induction Xs) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2337 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 2338 | lemma count_le_replicate_mset_subset_eq: "n \<le> count M x \<longleftrightarrow> replicate_mset n x \<subseteq># M" | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 2339 | by (auto simp add: mset_subset_eqI) (metis count_replicate_mset subseteq_mset_def) | 
| 60804 | 2340 | |
| 66494 | 2341 | lemma replicate_count_mset_eq_filter_eq: "replicate (count (mset xs) k) k = filter (HOL.eq k) xs" | 
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 2342 | by (induct xs) auto | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 2343 | |
| 66494 | 2344 | lemma replicate_mset_eq_empty_iff [simp]: "replicate_mset n a = {#} \<longleftrightarrow> n = 0"
 | 
| 62366 | 2345 | by (induct n) simp_all | 
| 2346 | ||
| 2347 | lemma replicate_mset_eq_iff: | |
| 66494 | 2348 | "replicate_mset m a = replicate_mset n b \<longleftrightarrow> m = 0 \<and> n = 0 \<or> m = n \<and> a = b" | 
| 62366 | 2349 | by (auto simp add: multiset_eq_iff) | 
| 2350 | ||
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 2351 | lemma repeat_mset_cancel1: "repeat_mset a A = repeat_mset a B \<longleftrightarrow> A = B \<or> a = 0" | 
| 63849 
0dd6731060d7
delete looping simp rule
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63831diff
changeset | 2352 | by (auto simp: multiset_eq_iff) | 
| 
0dd6731060d7
delete looping simp rule
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63831diff
changeset | 2353 | |
| 63908 
ca41b6670904
support replicate_mset in multiset simproc
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63882diff
changeset | 2354 | lemma repeat_mset_cancel2: "repeat_mset a A = repeat_mset b A \<longleftrightarrow> a = b \<or> A = {#}"
 | 
| 63849 
0dd6731060d7
delete looping simp rule
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63831diff
changeset | 2355 | by (auto simp: multiset_eq_iff) | 
| 
0dd6731060d7
delete looping simp rule
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63831diff
changeset | 2356 | |
| 64077 | 2357 | lemma repeat_mset_eq_empty_iff: "repeat_mset n A = {#} \<longleftrightarrow> n = 0 \<or> A = {#}"
 | 
| 2358 | by (cases n) auto | |
| 2359 | ||
| 63924 | 2360 | lemma image_replicate_mset [simp]: | 
| 2361 | "image_mset f (replicate_mset n a) = replicate_mset n (f a)" | |
| 2362 | by (induct n) simp_all | |
| 2363 | ||
| 67051 | 2364 | lemma replicate_mset_msubseteq_iff: | 
| 2365 | "replicate_mset m a \<subseteq># replicate_mset n b \<longleftrightarrow> m = 0 \<or> a = b \<and> m \<le> n" | |
| 2366 | by (cases m) | |
| 68406 | 2367 | (auto simp: insert_subset_eq_iff simp flip: count_le_replicate_mset_subset_eq) | 
| 67051 | 2368 | |
| 2369 | lemma msubseteq_replicate_msetE: | |
| 2370 | assumes "A \<subseteq># replicate_mset n a" | |
| 2371 | obtains m where "m \<le> n" and "A = replicate_mset m a" | |
| 2372 | proof (cases "n = 0") | |
| 2373 | case True | |
| 2374 | with assms that show thesis | |
| 2375 | by simp | |
| 2376 | next | |
| 2377 | case False | |
| 2378 | from assms have "set_mset A \<subseteq> set_mset (replicate_mset n a)" | |
| 2379 | by (rule set_mset_mono) | |
| 2380 |   with False have "set_mset A \<subseteq> {a}"
 | |
| 2381 | by simp | |
| 2382 | then have "\<exists>m. A = replicate_mset m a" | |
| 2383 | proof (induction A) | |
| 2384 | case empty | |
| 2385 | then show ?case | |
| 2386 | by simp | |
| 2387 | next | |
| 2388 | case (add b A) | |
| 2389 | then obtain m where "A = replicate_mset m a" | |
| 2390 | by auto | |
| 2391 | with add.prems show ?case | |
| 2392 | by (auto intro: exI [of _ "Suc m"]) | |
| 2393 | qed | |
| 2394 | then obtain m where A: "A = replicate_mset m a" .. | |
| 2395 | with assms have "m \<le> n" | |
| 2396 | by (auto simp add: replicate_mset_msubseteq_iff) | |
| 2397 | then show thesis using A .. | |
| 2398 | qed | |
| 2399 | ||
| 79800 | 2400 | lemma count_image_mset_lt_imp_lt_raw: | 
| 2401 | assumes | |
| 2402 | "finite A" and | |
| 2403 | "A = set_mset M \<union> set_mset N" and | |
| 2404 | "count (image_mset f M) b < count (image_mset f N) b" | |
| 2405 | shows "\<exists>x. f x = b \<and> count M x < count N x" | |
| 2406 | using assms | |
| 2407 | proof (induct A arbitrary: M N b rule: finite_induct) | |
| 2408 | case (insert x F) | |
| 2409 | note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and | |
| 2410 | cnt_b = this(5) | |
| 2411 | ||
| 2412 |   let ?Ma = "{#y \<in># M. y \<noteq> x#}"
 | |
| 2413 |   let ?Mb = "{#y \<in># M. y = x#}"
 | |
| 2414 |   let ?Na = "{#y \<in># N. y \<noteq> x#}"
 | |
| 2415 |   let ?Nb = "{#y \<in># N. y = x#}"
 | |
| 2416 | ||
| 2417 | have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" | |
| 2418 | using multiset_partition by blast+ | |
| 2419 | ||
| 2420 | have f_eq_ma_na: "F = set_mset ?Ma \<union> set_mset ?Na" | |
| 2421 | using x_f_eq_m_n x_ni_f by auto | |
| 2422 | ||
| 2423 | show ?case | |
| 2424 | proof (cases "count (image_mset f ?Ma) b < count (image_mset f ?Na) b") | |
| 2425 | case cnt_ba: True | |
| 2426 | obtain xa where "f xa = b" and "count ?Ma xa < count ?Na xa" | |
| 2427 | using ih[OF f_eq_ma_na cnt_ba] by blast | |
| 2428 | thus ?thesis | |
| 2429 | by (metis count_filter_mset not_less0) | |
| 2430 | next | |
| 2431 | case cnt_ba: False | |
| 2432 | have fx_eq_b: "f x = b" | |
| 2433 | using cnt_b cnt_ba | |
| 2434 | by (subst (asm) m_part, subst (asm) n_part, | |
| 2435 | auto simp: filter_eq_replicate_mset split: if_splits) | |
| 2436 | moreover have "count M x < count N x" | |
| 2437 | using cnt_b cnt_ba | |
| 2438 | by (subst (asm) m_part, subst (asm) n_part, | |
| 2439 | auto simp: filter_eq_replicate_mset split: if_splits) | |
| 2440 | ultimately show ?thesis | |
| 2441 | by blast | |
| 2442 | qed | |
| 2443 | qed auto | |
| 2444 | ||
| 2445 | lemma count_image_mset_lt_imp_lt: | |
| 2446 | assumes cnt_b: "count (image_mset f M) b < count (image_mset f N) b" | |
| 2447 | shows "\<exists>x. f x = b \<and> count M x < count N x" | |
| 2448 | by (rule count_image_mset_lt_imp_lt_raw[of "set_mset M \<union> set_mset N", OF _ refl cnt_b]) auto | |
| 2449 | ||
| 2450 | lemma count_image_mset_le_imp_lt_raw: | |
| 2451 | assumes | |
| 2452 | "finite A" and | |
| 2453 | "A = set_mset M \<union> set_mset N" and | |
| 2454 | "count (image_mset f M) (f a) + count N a < count (image_mset f N) (f a) + count M a" | |
| 2455 | shows "\<exists>b. f b = f a \<and> count M b < count N b" | |
| 2456 | using assms | |
| 2457 | proof (induct A arbitrary: M N rule: finite_induct) | |
| 2458 | case (insert x F) | |
| 2459 | note fin = this(1) and x_ni_f = this(2) and ih = this(3) and x_f_eq_m_n = this(4) and | |
| 2460 | cnt_lt = this(5) | |
| 2461 | ||
| 2462 |   let ?Ma = "{#y \<in># M. y \<noteq> x#}"
 | |
| 2463 |   let ?Mb = "{#y \<in># M. y = x#}"
 | |
| 2464 |   let ?Na = "{#y \<in># N. y \<noteq> x#}"
 | |
| 2465 |   let ?Nb = "{#y \<in># N. y = x#}"
 | |
| 2466 | ||
| 2467 | have m_part: "M = ?Mb + ?Ma" and n_part: "N = ?Nb + ?Na" | |
| 2468 | using multiset_partition by blast+ | |
| 2469 | ||
| 2470 | have f_eq_ma_na: "F = set_mset ?Ma \<union> set_mset ?Na" | |
| 2471 | using x_f_eq_m_n x_ni_f by auto | |
| 2472 | ||
| 2473 | show ?case | |
| 2474 | proof (cases "f x = f a") | |
| 2475 | case fx_ne_fa: False | |
| 2476 | ||
| 2477 | have cnt_fma_fa: "count (image_mset f ?Ma) (f a) = count (image_mset f M) (f a)" | |
| 2478 | using fx_ne_fa by (subst (2) m_part) (auto simp: filter_eq_replicate_mset) | |
| 2479 | have cnt_fna_fa: "count (image_mset f ?Na) (f a) = count (image_mset f N) (f a)" | |
| 2480 | using fx_ne_fa by (subst (2) n_part) (auto simp: filter_eq_replicate_mset) | |
| 2481 | have cnt_ma_a: "count ?Ma a = count M a" | |
| 2482 | using fx_ne_fa by (subst (2) m_part) (auto simp: filter_eq_replicate_mset) | |
| 2483 | have cnt_na_a: "count ?Na a = count N a" | |
| 2484 | using fx_ne_fa by (subst (2) n_part) (auto simp: filter_eq_replicate_mset) | |
| 2485 | ||
| 2486 | obtain b where fb_eq_fa: "f b = f a" and cnt_b: "count ?Ma b < count ?Na b" | |
| 2487 | using ih[OF f_eq_ma_na] cnt_lt unfolding cnt_fma_fa cnt_fna_fa cnt_ma_a cnt_na_a by blast | |
| 2488 | have fx_ne_fb: "f x \<noteq> f b" | |
| 2489 | using fb_eq_fa fx_ne_fa by simp | |
| 2490 | ||
| 2491 | have cnt_ma_b: "count ?Ma b = count M b" | |
| 2492 | using fx_ne_fb by (subst (2) m_part) auto | |
| 2493 | have cnt_na_b: "count ?Na b = count N b" | |
| 2494 | using fx_ne_fb by (subst (2) n_part) auto | |
| 2495 | ||
| 2496 | show ?thesis | |
| 2497 | using fb_eq_fa cnt_b unfolding cnt_ma_b cnt_na_b by blast | |
| 2498 | next | |
| 2499 | case fx_eq_fa: True | |
| 2500 | show ?thesis | |
| 2501 | proof (cases "x = a") | |
| 2502 | case x_eq_a: True | |
| 2503 | have "count (image_mset f ?Ma) (f a) + count ?Na a | |
| 2504 | < count (image_mset f ?Na) (f a) + count ?Ma a" | |
| 2505 | using cnt_lt x_eq_a by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, | |
| 2506 | auto simp: filter_eq_replicate_mset) | |
| 2507 | thus ?thesis | |
| 2508 | using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) | |
| 2509 | next | |
| 2510 | case x_ne_a: False | |
| 2511 | show ?thesis | |
| 2512 | proof (cases "count M x < count N x") | |
| 2513 | case True | |
| 2514 | thus ?thesis | |
| 2515 | using fx_eq_fa by blast | |
| 2516 | next | |
| 2517 | case False | |
| 2518 | hence cnt_x: "count M x \<ge> count N x" | |
| 2519 | by fastforce | |
| 2520 | have "count M x + count (image_mset f ?Ma) (f a) + count ?Na a | |
| 2521 | < count N x + count (image_mset f ?Na) (f a) + count ?Ma a" | |
| 2522 | using cnt_lt x_ne_a fx_eq_fa by (subst (asm) (1 2) m_part, subst (asm) (1 2) n_part, | |
| 2523 | auto simp: filter_eq_replicate_mset) | |
| 2524 | hence "count (image_mset f ?Ma) (f a) + count ?Na a | |
| 2525 | < count (image_mset f ?Na) (f a) + count ?Ma a" | |
| 2526 | using cnt_x by linarith | |
| 2527 | thus ?thesis | |
| 2528 | using ih[OF f_eq_ma_na] by (metis count_filter_mset nat_neq_iff) | |
| 2529 | qed | |
| 2530 | qed | |
| 2531 | qed | |
| 2532 | qed auto | |
| 2533 | ||
| 2534 | lemma count_image_mset_le_imp_lt: | |
| 2535 | assumes | |
| 2536 | "count (image_mset f M) (f a) \<le> count (image_mset f N) (f a)" and | |
| 2537 | "count M a > count N a" | |
| 2538 | shows "\<exists>b. f b = f a \<and> count M b < count N b" | |
| 2539 | using assms by (auto intro: count_image_mset_le_imp_lt_raw[of "set_mset M \<union> set_mset N"]) | |
| 2540 | ||
| 2541 | lemma size_filter_unsat_elem: | |
| 2542 | assumes "x \<in># M" and "\<not> P x" | |
| 2543 |   shows "size {#x \<in># M. P x#} < size M"
 | |
| 2544 | proof - | |
| 2545 | have "size (filter_mset P M) \<noteq> size M" | |
| 80095 | 2546 | using assms | 
| 2547 | by (metis dual_order.strict_iff_order filter_mset_eq_conv mset_subset_size subset_mset.nless_le) | |
| 79800 | 2548 | then show ?thesis | 
| 2549 | by (meson leD nat_neq_iff size_filter_mset_lesseq) | |
| 2550 | qed | |
| 2551 | ||
| 2552 | lemma size_filter_ne_elem: "x \<in># M \<Longrightarrow> size {#y \<in># M. y \<noteq> x#} < size M"
 | |
| 2553 | by (simp add: size_filter_unsat_elem[of x M "\<lambda>y. y \<noteq> x"]) | |
| 2554 | ||
| 2555 | lemma size_eq_ex_count_lt: | |
| 2556 | assumes | |
| 2557 | sz_m_eq_n: "size M = size N" and | |
| 2558 | m_eq_n: "M \<noteq> N" | |
| 2559 | shows "\<exists>x. count M x < count N x" | |
| 2560 | proof - | |
| 2561 | obtain x where "count M x \<noteq> count N x" | |
| 2562 | using m_eq_n by (meson multiset_eqI) | |
| 2563 | moreover | |
| 2564 |   {
 | |
| 2565 | assume "count M x < count N x" | |
| 2566 | hence ?thesis | |
| 2567 | by blast | |
| 2568 | } | |
| 2569 | moreover | |
| 2570 |   {
 | |
| 2571 | assume cnt_x: "count M x > count N x" | |
| 2572 | ||
| 2573 |     have "size {#y \<in># M. y = x#} + size {#y \<in># M. y \<noteq> x#} =
 | |
| 2574 |       size {#y \<in># N. y = x#} + size {#y \<in># N. y \<noteq> x#}"
 | |
| 2575 | using sz_m_eq_n multiset_partition by (metis size_union) | |
| 2576 |     hence sz_m_minus_x: "size {#y \<in># M. y \<noteq> x#} < size {#y \<in># N. y \<noteq> x#}"
 | |
| 2577 | using cnt_x by (simp add: filter_eq_replicate_mset) | |
| 2578 |     then obtain y where "count {#y \<in># M. y \<noteq> x#} y < count {#y \<in># N. y \<noteq> x#} y"
 | |
| 2579 | using size_lt_imp_ex_count_lt[OF sz_m_minus_x] by blast | |
| 2580 | hence "count M y < count N y" | |
| 2581 | by (metis count_filter_mset less_asym) | |
| 2582 | hence ?thesis | |
| 2583 | by blast | |
| 2584 | } | |
| 2585 | ultimately show ?thesis | |
| 2586 | by fastforce | |
| 2587 | qed | |
| 2588 | ||
| 60804 | 2589 | |
| 60500 | 2590 | subsection \<open>Big operators\<close> | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2591 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2592 | locale comm_monoid_mset = comm_monoid | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2593 | begin | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2594 | |
| 64075 | 2595 | interpretation comp_fun_commute f | 
| 2596 | by standard (simp add: fun_eq_iff left_commute) | |
| 2597 | ||
| 2598 | interpretation comp?: comp_fun_commute "f \<circ> g" | |
| 2599 | by (fact comp_comp_fun_commute) | |
| 2600 | ||
| 2601 | context | |
| 2602 | begin | |
| 2603 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2604 | definition F :: "'a multiset \<Rightarrow> 'a" | 
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63195diff
changeset | 2605 | where eq_fold: "F M = fold_mset f \<^bold>1 M" | 
| 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63195diff
changeset | 2606 | |
| 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63195diff
changeset | 2607 | lemma empty [simp]: "F {#} = \<^bold>1"
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2608 | by (simp add: eq_fold) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2609 | |
| 60678 | 2610 | lemma singleton [simp]: "F {#x#} = x"
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2611 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2612 | interpret comp_fun_commute | 
| 60678 | 2613 | by standard (simp add: fun_eq_iff left_commute) | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2614 | show ?thesis by (simp add: eq_fold) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2615 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2616 | |
| 63290 
9ac558ab0906
boldify syntax in abstract algebraic structures, to avoid clashes with concrete syntax in corresponding type classes
 haftmann parents: 
63195diff
changeset | 2617 | lemma union [simp]: "F (M + N) = F M \<^bold>* F N" | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2618 | proof - | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2619 | interpret comp_fun_commute f | 
| 60678 | 2620 | by standard (simp add: fun_eq_iff left_commute) | 
| 2621 | show ?thesis | |
| 2622 | by (induct N) (simp_all add: left_commute eq_fold) | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2623 | qed | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2624 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2625 | lemma add_mset [simp]: "F (add_mset x N) = x \<^bold>* F N" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2626 | unfolding add_mset_add_single[of x N] union by (simp add: ac_simps) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2627 | |
| 64075 | 2628 | lemma insert [simp]: | 
| 2629 | shows "F (image_mset g (add_mset x A)) = g x \<^bold>* F (image_mset g A)" | |
| 2630 | by (simp add: eq_fold) | |
| 2631 | ||
| 2632 | lemma remove: | |
| 2633 | assumes "x \<in># A" | |
| 2634 |   shows "F A = x \<^bold>* F (A - {#x#})"
 | |
| 2635 | using multi_member_split[OF assms] by auto | |
| 2636 | ||
| 2637 | lemma neutral: | |
| 2638 | "\<forall>x\<in>#A. x = \<^bold>1 \<Longrightarrow> F A = \<^bold>1" | |
| 2639 | by (induct A) simp_all | |
| 2640 | ||
| 2641 | lemma neutral_const [simp]: | |
| 2642 | "F (image_mset (\<lambda>_. \<^bold>1) A) = \<^bold>1" | |
| 2643 | by (simp add: neutral) | |
| 2644 | ||
| 2645 | private lemma F_image_mset_product: | |
| 2646 |   "F {#g x j \<^bold>* F {#g i j. i \<in># A#}. j \<in># B#} =
 | |
| 2647 |     F (image_mset (g x) B) \<^bold>* F {#F {#g i j. i \<in># A#}. j \<in># B#}"
 | |
| 2648 | by (induction B) (simp_all add: left_commute semigroup.assoc semigroup_axioms) | |
| 2649 | ||
| 68938 | 2650 | lemma swap: | 
| 64075 | 2651 | "F (image_mset (\<lambda>i. F (image_mset (g i) B)) A) = | 
| 2652 | F (image_mset (\<lambda>j. F (image_mset (\<lambda>i. g i j) A)) B)" | |
| 2653 | apply (induction A, simp) | |
| 2654 | apply (induction B, auto simp add: F_image_mset_product ac_simps) | |
| 2655 | done | |
| 2656 | ||
| 2657 | lemma distrib: "F (image_mset (\<lambda>x. g x \<^bold>* h x) A) = F (image_mset g A) \<^bold>* F (image_mset h A)" | |
| 2658 | by (induction A) (auto simp: ac_simps) | |
| 2659 | ||
| 2660 | lemma union_disjoint: | |
| 2661 |   "A \<inter># B = {#} \<Longrightarrow> F (image_mset g (A \<union># B)) = F (image_mset g A) \<^bold>* F (image_mset g B)"
 | |
| 2662 | by (induction A) (auto simp: ac_simps) | |
| 2663 | ||
| 2664 | end | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2665 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2666 | |
| 67398 | 2667 | lemma comp_fun_commute_plus_mset[simp]: "comp_fun_commute ((+) :: 'a multiset \<Rightarrow> _ \<Rightarrow> _)" | 
| 60678 | 2668 | by standard (simp add: add_ac comp_def) | 
| 59813 | 2669 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2670 | declare comp_fun_commute.fold_mset_add_mset[OF comp_fun_commute_plus_mset, simp] | 
| 59813 | 2671 | |
| 67398 | 2672 | lemma in_mset_fold_plus_iff[iff]: "x \<in># fold_mset (+) M NN \<longleftrightarrow> x \<in># M \<or> (\<exists>N. N \<in># NN \<and> x \<in># N)" | 
| 59813 | 2673 | by (induct NN) auto | 
| 2674 | ||
| 54868 | 2675 | context comm_monoid_add | 
| 2676 | begin | |
| 2677 | ||
| 63830 | 2678 | sublocale sum_mset: comm_monoid_mset plus 0 | 
| 2679 | defines sum_mset = sum_mset.F .. | |
| 2680 | ||
| 64267 | 2681 | lemma sum_unfold_sum_mset: | 
| 2682 | "sum f A = sum_mset (image_mset f (mset_set A))" | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2683 | by (cases "finite A") (induct A rule: finite_induct, simp_all) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2684 | |
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2685 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2686 | |
| 73047 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
72607diff
changeset | 2687 | notation sum_mset ("\<Sum>\<^sub>#")
 | 
| 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
72607diff
changeset | 2688 | |
| 62366 | 2689 | syntax (ASCII) | 
| 63830 | 2690 |   "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3SUM _:#_. _)" [0, 51, 10] 10)
 | 
| 62366 | 2691 | syntax | 
| 63830 | 2692 |   "_sum_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_add"  ("(3\<Sum>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 80768 | 2693 | syntax_consts | 
| 2694 | "_sum_mset_image" \<rightleftharpoons> sum_mset | |
| 62366 | 2695 | translations | 
| 63830 | 2696 | "\<Sum>i \<in># A. b" \<rightleftharpoons> "CONST sum_mset (CONST image_mset (\<lambda>i. b) A)" | 
| 59949 | 2697 | |
| 66938 | 2698 | context comm_monoid_add | 
| 2699 | begin | |
| 2700 | ||
| 2701 | lemma sum_mset_sum_list: | |
| 2702 | "sum_mset (mset xs) = sum_list xs" | |
| 2703 | by (induction xs) auto | |
| 2704 | ||
| 2705 | end | |
| 2706 | ||
| 2707 | context canonically_ordered_monoid_add | |
| 2708 | begin | |
| 2709 | ||
| 2710 | lemma sum_mset_0_iff [simp]: | |
| 2711 | "sum_mset M = 0 \<longleftrightarrow> (\<forall>x \<in> set_mset M. x = 0)" | |
| 2712 | by (induction M) auto | |
| 2713 | ||
| 2714 | end | |
| 2715 | ||
| 2716 | context ordered_comm_monoid_add | |
| 2717 | begin | |
| 2718 | ||
| 2719 | lemma sum_mset_mono: | |
| 2720 | "sum_mset (image_mset f K) \<le> sum_mset (image_mset g K)" | |
| 2721 | if "\<And>i. i \<in># K \<Longrightarrow> f i \<le> g i" | |
| 2722 | using that by (induction K) (simp_all add: add_mono) | |
| 2723 | ||
| 2724 | end | |
| 2725 | ||
| 73470 | 2726 | context cancel_comm_monoid_add | 
| 66938 | 2727 | begin | 
| 2728 | ||
| 2729 | lemma sum_mset_diff: | |
| 2730 | "sum_mset (M - N) = sum_mset M - sum_mset N" if "N \<subseteq># M" for M N :: "'a multiset" | |
| 2731 | using that by (auto simp add: subset_mset.le_iff_add) | |
| 2732 | ||
| 2733 | end | |
| 2734 | ||
| 2735 | context semiring_0 | |
| 2736 | begin | |
| 2737 | ||
| 63860 | 2738 | lemma sum_mset_distrib_left: | 
| 66938 | 2739 | "c * (\<Sum>x \<in># M. f x) = (\<Sum>x \<in># M. c * f(x))" | 
| 2740 | by (induction M) (simp_all add: algebra_simps) | |
| 63860 | 2741 | |
| 64075 | 2742 | lemma sum_mset_distrib_right: | 
| 66938 | 2743 | "(\<Sum>x \<in># M. f x) * c = (\<Sum>x \<in># M. f x * c)" | 
| 2744 | by (induction M) (simp_all add: algebra_simps) | |
| 2745 | ||
| 2746 | end | |
| 2747 | ||
| 2748 | lemma sum_mset_product: | |
| 2749 |   fixes f :: "'a::{comm_monoid_add,times} \<Rightarrow> 'b::semiring_0"
 | |
| 2750 | shows "(\<Sum>i \<in># A. f i) * (\<Sum>i \<in># B. g i) = (\<Sum>i\<in>#A. \<Sum>j\<in>#B. f i * g j)" | |
| 68938 | 2751 | by (subst sum_mset.swap) (simp add: sum_mset_distrib_left sum_mset_distrib_right) | 
| 66938 | 2752 | |
| 2753 | context semiring_1 | |
| 2754 | begin | |
| 2755 | ||
| 2756 | lemma sum_mset_replicate_mset [simp]: | |
| 2757 | "sum_mset (replicate_mset n a) = of_nat n * a" | |
| 2758 | by (induction n) (simp_all add: algebra_simps) | |
| 2759 | ||
| 2760 | lemma sum_mset_delta: | |
| 2761 | "sum_mset (image_mset (\<lambda>x. if x = y then c else 0) A) = c * of_nat (count A y)" | |
| 2762 | by (induction A) (simp_all add: algebra_simps) | |
| 2763 | ||
| 2764 | lemma sum_mset_delta': | |
| 2765 | "sum_mset (image_mset (\<lambda>x. if y = x then c else 0) A) = c * of_nat (count A y)" | |
| 2766 | by (induction A) (simp_all add: algebra_simps) | |
| 2767 | ||
| 2768 | end | |
| 2769 | ||
| 2770 | lemma of_nat_sum_mset [simp]: | |
| 2771 | "of_nat (sum_mset A) = sum_mset (image_mset of_nat A)" | |
| 2772 | by (induction A) auto | |
| 2773 | ||
| 2774 | lemma size_eq_sum_mset: | |
| 2775 | "size M = (\<Sum>a\<in>#M. 1)" | |
| 2776 | using image_mset_const_eq [of "1::nat" M] by simp | |
| 2777 | ||
| 2778 | lemma size_mset_set [simp]: | |
| 2779 | "size (mset_set A) = card A" | |
| 2780 | by (simp only: size_eq_sum_mset card_eq_sum sum_unfold_sum_mset) | |
| 64075 | 2781 | |
| 2782 | lemma sum_mset_constant [simp]: | |
| 2783 | fixes y :: "'b::semiring_1" | |
| 2784 | shows \<open>(\<Sum>x\<in>#A. y) = of_nat (size A) * y\<close> | |
| 2785 | by (induction A) (auto simp: algebra_simps) | |
| 2786 | ||
| 73047 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
72607diff
changeset | 2787 | lemma set_mset_Union_mset[simp]: "set_mset (\<Sum>\<^sub># MM) = (\<Union>M \<in> set_mset MM. set_mset M)" | 
| 59813 | 2788 | by (induct MM) auto | 
| 2789 | ||
| 73047 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
72607diff
changeset | 2790 | lemma in_Union_mset_iff[iff]: "x \<in># \<Sum>\<^sub># MM \<longleftrightarrow> (\<exists>M. M \<in># MM \<and> x \<in># M)" | 
| 59813 | 2791 | by (induct MM) auto | 
| 2792 | ||
| 64267 | 2793 | lemma count_sum: | 
| 2794 | "count (sum f A) x = sum (\<lambda>a. count (f a) x) A" | |
| 62366 | 2795 | by (induct A rule: infinite_finite_induct) simp_all | 
| 2796 | ||
| 64267 | 2797 | lemma sum_eq_empty_iff: | 
| 62366 | 2798 | assumes "finite A" | 
| 64267 | 2799 |   shows "sum f A = {#} \<longleftrightarrow> (\<forall>a\<in>A. f a = {#})"
 | 
| 62366 | 2800 | using assms by induct simp_all | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2801 | |
| 73047 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
72607diff
changeset | 2802 | lemma Union_mset_empty_conv[simp]: "\<Sum>\<^sub># M = {#} \<longleftrightarrow> (\<forall>i\<in>#M. i = {#})"
 | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 2803 | by (induction M) auto | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 2804 | |
| 73047 
ab9e27da0e85
HOL-Library: Changed notation for sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
72607diff
changeset | 2805 | lemma Union_image_single_mset[simp]: "\<Sum>\<^sub># (image_mset (\<lambda>x. {#x#}) m) = m"
 | 
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2806 | by(induction m) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2807 | |
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2808 | lemma size_multiset_sum_mset [simp]: "size (\<Sum>X\<in>#A. X :: 'a multiset) = (\<Sum>X\<in>#A. size X)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2809 | by (induction A) auto | 
| 67656 | 2810 | |
| 54868 | 2811 | context comm_monoid_mult | 
| 2812 | begin | |
| 2813 | ||
| 63830 | 2814 | sublocale prod_mset: comm_monoid_mset times 1 | 
| 2815 | defines prod_mset = prod_mset.F .. | |
| 2816 | ||
| 2817 | lemma prod_mset_empty: | |
| 2818 |   "prod_mset {#} = 1"
 | |
| 2819 | by (fact prod_mset.empty) | |
| 2820 | ||
| 2821 | lemma prod_mset_singleton: | |
| 2822 |   "prod_mset {#x#} = x"
 | |
| 2823 | by (fact prod_mset.singleton) | |
| 2824 | ||
| 2825 | lemma prod_mset_Un: | |
| 2826 | "prod_mset (A + B) = prod_mset A * prod_mset B" | |
| 2827 | by (fact prod_mset.union) | |
| 2828 | ||
| 66938 | 2829 | lemma prod_mset_prod_list: | 
| 2830 | "prod_mset (mset xs) = prod_list xs" | |
| 2831 | by (induct xs) auto | |
| 2832 | ||
| 63830 | 2833 | lemma prod_mset_replicate_mset [simp]: | 
| 2834 | "prod_mset (replicate_mset n a) = a ^ n" | |
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2835 | by (induct n) simp_all | 
| 60804 | 2836 | |
| 64272 | 2837 | lemma prod_unfold_prod_mset: | 
| 2838 | "prod f A = prod_mset (image_mset f (mset_set A))" | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2839 | by (cases "finite A") (induct A rule: finite_induct, simp_all) | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2840 | |
| 63830 | 2841 | lemma prod_mset_multiplicity: | 
| 64272 | 2842 | "prod_mset M = prod (\<lambda>x. x ^ count M x) (set_mset M)" | 
| 2843 | by (simp add: fold_mset_def prod.eq_fold prod_mset.eq_fold funpow_times_power comp_def) | |
| 63830 | 2844 | |
| 2845 | lemma prod_mset_delta: "prod_mset (image_mset (\<lambda>x. if x = y then c else 1) A) = c ^ count A y" | |
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2846 | by (induction A) simp_all | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63524diff
changeset | 2847 | |
| 63830 | 2848 | lemma prod_mset_delta': "prod_mset (image_mset (\<lambda>x. if y = x then c else 1) A) = c ^ count A y" | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2849 | by (induction A) simp_all | 
| 63534 
523b488b15c9
Overhaul of prime/multiplicity/prime_factors
 eberlm <eberlm@in.tum.de> parents: 
63524diff
changeset | 2850 | |
| 66938 | 2851 | lemma prod_mset_subset_imp_dvd: | 
| 2852 | assumes "A \<subseteq># B" | |
| 2853 | shows "prod_mset A dvd prod_mset B" | |
| 2854 | proof - | |
| 2855 | from assms have "B = (B - A) + A" by (simp add: subset_mset.diff_add) | |
| 2856 | also have "prod_mset \<dots> = prod_mset (B - A) * prod_mset A" by simp | |
| 2857 | also have "prod_mset A dvd \<dots>" by simp | |
| 2858 | finally show ?thesis . | |
| 2859 | qed | |
| 2860 | ||
| 2861 | lemma dvd_prod_mset: | |
| 2862 | assumes "x \<in># A" | |
| 2863 | shows "x dvd prod_mset A" | |
| 2864 |   using assms prod_mset_subset_imp_dvd [of "{#x#}" A] by simp
 | |
| 2865 | ||
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2866 | end | 
| 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2867 | |
| 73052 
c03a148110cc
HOL-Library.Multiset: new notation for prod_mset, consistent with sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
73047diff
changeset | 2868 | notation prod_mset ("\<Prod>\<^sub>#")
 | 
| 
c03a148110cc
HOL-Library.Multiset: new notation for prod_mset, consistent with sum_mset
 Manuel Eberl <eberlm@in.tum.de> parents: 
73047diff
changeset | 2869 | |
| 61955 
e96292f32c3c
former "xsymbols" syntax is used by default, and ASCII replacement syntax with print mode "ASCII";
 wenzelm parents: 
61890diff
changeset | 2870 | syntax (ASCII) | 
| 63830 | 2871 |   "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3PROD _:#_. _)" [0, 51, 10] 10)
 | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2872 | syntax | 
| 63830 | 2873 |   "_prod_mset_image" :: "pttrn \<Rightarrow> 'b set \<Rightarrow> 'a \<Rightarrow> 'a::comm_monoid_mult"  ("(3\<Prod>_\<in>#_. _)" [0, 51, 10] 10)
 | 
| 80768 | 2874 | syntax_consts | 
| 2875 | "_prod_mset_image" \<rightleftharpoons> prod_mset | |
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2876 | translations | 
| 63830 | 2877 | "\<Prod>i \<in># A. b" \<rightleftharpoons> "CONST prod_mset (CONST image_mset (\<lambda>i. b) A)" | 
| 2878 | ||
| 64591 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64587diff
changeset | 2879 | lemma prod_mset_constant [simp]: "(\<Prod>_\<in>#A. c) = c ^ size A" | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64587diff
changeset | 2880 | by (simp add: image_mset_const_eq) | 
| 
240a39af9ec4
restructured matter on polynomials and normalized fractions
 haftmann parents: 
64587diff
changeset | 2881 | |
| 63830 | 2882 | lemma (in semidom) prod_mset_zero_iff [iff]: | 
| 2883 | "prod_mset A = 0 \<longleftrightarrow> 0 \<in># A" | |
| 62366 | 2884 | by (induct A) auto | 
| 2885 | ||
| 63830 | 2886 | lemma (in semidom_divide) prod_mset_diff: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2887 | assumes "B \<subseteq># A" and "0 \<notin># B" | 
| 63830 | 2888 | shows "prod_mset (A - B) = prod_mset A div prod_mset B" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2889 | proof - | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2890 | from assms obtain C where "A = B + C" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2891 | by (metis subset_mset.add_diff_inverse) | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2892 | with assms show ?thesis by simp | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2893 | qed | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2894 | |
| 63830 | 2895 | lemma (in semidom_divide) prod_mset_minus: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2896 | assumes "a \<in># A" and "a \<noteq> 0" | 
| 63830 | 2897 |   shows "prod_mset (A - {#a#}) = prod_mset A div a"
 | 
| 2898 |   using assms prod_mset_diff [of "{#a#}" A] by auto
 | |
| 2899 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2900 | lemma (in normalization_semidom) normalize_prod_mset_normalize: | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2901 | "normalize (prod_mset (image_mset normalize A)) = normalize (prod_mset A)" | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2902 | proof (induction A) | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2903 | case (add x A) | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2904 | have "normalize (prod_mset (image_mset normalize (add_mset x A))) = | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2905 | normalize (x * normalize (prod_mset (image_mset normalize A)))" | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2906 | by simp | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2907 | also note add.IH | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2908 | finally show ?case by simp | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2909 | qed auto | 
| 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2910 | |
| 63924 | 2911 | lemma (in algebraic_semidom) is_unit_prod_mset_iff: | 
| 2912 | "is_unit (prod_mset A) \<longleftrightarrow> (\<forall>x \<in># A. is_unit x)" | |
| 2913 | by (induct A) (auto simp: is_unit_mult_iff) | |
| 2914 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2915 | lemma (in normalization_semidom_multiplicative) normalize_prod_mset: | 
| 63924 | 2916 | "normalize (prod_mset A) = prod_mset (image_mset normalize A)" | 
| 2917 | by (induct A) (simp_all add: normalize_mult) | |
| 2918 | ||
| 71398 
e0237f2eb49d
Removed multiplicativity assumption from normalization_semidom
 Manuel Eberl <eberlm@in.tum.de> parents: 
69895diff
changeset | 2919 | lemma (in normalization_semidom_multiplicative) normalized_prod_msetI: | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2920 | assumes "\<And>a. a \<in># A \<Longrightarrow> normalize a = a" | 
| 63830 | 2921 | shows "normalize (prod_mset A) = prod_mset A" | 
| 63924 | 2922 | proof - | 
| 2923 | from assms have "image_mset normalize A = A" | |
| 2924 | by (induct A) simp_all | |
| 2925 | then show ?thesis by (simp add: normalize_prod_mset) | |
| 2926 | qed | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2927 | |
| 80061 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2928 | lemma image_prod_mset_multiplicity: | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2929 | "prod_mset (image_mset f M) = prod (\<lambda>x. f x ^ count M x) (set_mset M)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2930 | proof (induction M) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2931 | case (add x M) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2932 | show ?case | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2933 | proof (cases "x \<in> set_mset M") | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2934 | case True | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2935 | have "(\<Prod>y\<in>set_mset (add_mset x M). f y ^ count (add_mset x M) y) = | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2936 | (\<Prod>y\<in>set_mset M. (if y = x then f x else 1) * f y ^ count M y)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2937 | using True add by (intro prod.cong) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2938 | also have "\<dots> = f x * (\<Prod>y\<in>set_mset M. f y ^ count M y)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2939 | using True by (subst prod.distrib) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2940 | also note add.IH [symmetric] | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2941 | finally show ?thesis using True by simp | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2942 | next | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2943 | case False | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2944 | hence "(\<Prod>y\<in>set_mset (add_mset x M). f y ^ count (add_mset x M) y) = | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2945 | f x * (\<Prod>y\<in>set_mset M. f y ^ count (add_mset x M) y)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2946 | by (auto simp: not_in_iff) | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2947 | also have "(\<Prod>y\<in>set_mset M. f y ^ count (add_mset x M) y) = | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2948 | (\<Prod>y\<in>set_mset M. f y ^ count M y)" | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2949 | using False by (intro prod.cong) auto | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2950 | also note add.IH [symmetric] | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2951 | finally show ?thesis by simp | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2952 | qed | 
| 
4c1347e172b1
moved over material from AFP; most importantly on algebraic numbers and algebraically closed fields
 Manuel Eberl <eberlm@in.tum.de> parents: 
79971diff
changeset | 2953 | qed auto | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2954 | |
| 73301 | 2955 | subsection \<open>Multiset as order-ignorant lists\<close> | 
| 51548 
757fa47af981
centralized various multiset operations in theory multiset;
 haftmann parents: 
51161diff
changeset | 2956 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2957 | context linorder | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2958 | begin | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2959 | |
| 60515 | 2960 | lemma mset_insort [simp]: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2961 | "mset (insort_key k x xs) = add_mset x (mset xs)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 2962 | by (induct xs) simp_all | 
| 58425 | 2963 | |
| 60515 | 2964 | lemma mset_sort [simp]: | 
| 2965 | "mset (sort_key k xs) = mset xs" | |
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 2966 | by (induct xs) simp_all | 
| 37107 | 2967 | |
| 60500 | 2968 | 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 | 2969 | This lemma shows which properties suffice to show that a function | 
| 61585 | 2970 | \<open>f\<close> with \<open>f xs = ys\<close> behaves like sort. | 
| 60500 | 2971 | \<close> | 
| 37074 | 2972 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2973 | lemma properties_for_sort_key: | 
| 60515 | 2974 | assumes "mset ys = mset xs" | 
| 60606 | 2975 | and "\<And>k. k \<in> set ys \<Longrightarrow> filter (\<lambda>x. f k = f x) ys = filter (\<lambda>x. f k = f x) xs" | 
| 2976 | 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 | 2977 | shows "sort_key f xs = ys" | 
| 60606 | 2978 | using assms | 
| 46921 | 2979 | 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 | 2980 | 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 | 2981 | 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 | 2982 | case (Cons x xs) | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2983 | from Cons.prems(2) have | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 2984 | "\<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 | 2985 | by (simp add: filter_remove1) | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2986 | 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 | 2987 | 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 | 2988 | moreover from Cons.prems have "x \<in># mset ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2989 | by auto | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2990 | then have "x \<in> set ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 2991 | by simp | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2992 | 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 | 2993 | 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 | 2994 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2995 | lemma properties_for_sort: | 
| 60515 | 2996 | assumes multiset: "mset ys = mset xs" | 
| 60606 | 2997 | and "sorted ys" | 
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2998 | shows "sort xs = ys" | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 2999 | proof (rule properties_for_sort_key) | 
| 60515 | 3000 | from multiset show "mset ys = mset xs" . | 
| 60500 | 3001 | from \<open>sorted ys\<close> show "sorted (map (\<lambda>x. x) ys)" by simp | 
| 60678 | 3002 | from multiset have "length (filter (\<lambda>y. k = y) ys) = length (filter (\<lambda>x. k = x) xs)" for k | 
| 60515 | 3003 | by (rule mset_eq_length_filter) | 
| 60678 | 3004 | then have "replicate (length (filter (\<lambda>y. k = y) ys)) k = | 
| 3005 | 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 | 3006 | by simp | 
| 60678 | 3007 | 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 | 3008 | 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 | 3009 | qed | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 3010 | |
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3011 | lemma sort_key_inj_key_eq: | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3012 | assumes mset_equal: "mset xs = mset ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3013 | and "inj_on f (set xs)" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3014 | and "sorted (map f ys)" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3015 | shows "sort_key f xs = ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3016 | proof (rule properties_for_sort_key) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3017 | from mset_equal | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3018 | show "mset ys = mset xs" by simp | 
| 61188 | 3019 | from \<open>sorted (map f ys)\<close> | 
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3020 | show "sorted (map f ys)" . | 
| 68386 | 3021 | show "[x\<leftarrow>ys . f k = f x] = [x\<leftarrow>xs . f k = f x]" if "k \<in> set ys" for k | 
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3022 | proof - | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3023 | from mset_equal | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3024 | 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 | 3025 | with that have "insert k (set ys) = set ys" by auto | 
| 61188 | 3026 | 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 | 3027 | by (simp add: set_equal) | 
| 68386 | 3028 | from inj have "[x\<leftarrow>ys . f k = f x] = filter (HOL.eq k) ys" | 
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3029 | by (auto intro!: inj_on_filter_key_eq) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3030 | also have "\<dots> = replicate (count (mset ys) k) k" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3031 | by (simp add: replicate_count_mset_eq_filter_eq) | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3032 | also have "\<dots> = replicate (count (mset xs) k) k" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3033 | using mset_equal by simp | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3034 | also have "\<dots> = filter (HOL.eq k) xs" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3035 | by (simp add: replicate_count_mset_eq_filter_eq) | 
| 68386 | 3036 | also have "\<dots> = [x\<leftarrow>xs . f k = f x]" | 
| 61031 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3037 | 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 | 3038 | finally show ?thesis . | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3039 | qed | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3040 | qed | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3041 | |
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3042 | lemma sort_key_eq_sort_key: | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3043 | assumes "mset xs = mset ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3044 | and "inj_on f (set xs)" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3045 | shows "sort_key f xs = sort_key f ys" | 
| 
162bd20dae23
more lemmas on sorting and multisets (due to Thomas Sewell)
 haftmann parents: 
60804diff
changeset | 3046 | 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 | 3047 | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3048 | lemma sort_key_by_quicksort: | 
| 68386 | 3049 | "sort_key f xs = sort_key f [x\<leftarrow>xs. f x < f (xs ! (length xs div 2))] | 
| 3050 | @ [x\<leftarrow>xs. f x = f (xs ! (length xs div 2))] | |
| 3051 | @ sort_key f [x\<leftarrow>xs. f x > f (xs ! (length xs div 2))]" (is "sort_key f ?lhs = ?rhs") | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3052 | proof (rule properties_for_sort_key) | 
| 60515 | 3053 | show "mset ?rhs = mset ?lhs" | 
| 69442 | 3054 | by (rule multiset_eqI) auto | 
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3055 | show "sorted (map f ?rhs)" | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3056 | by (auto simp add: sorted_append intro: sorted_map_same) | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3057 | next | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 3058 | fix l | 
| 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 3059 | assume "l \<in> set ?rhs" | 
| 40346 | 3060 | let ?pivot = "f (xs ! (length xs div 2))" | 
| 3061 | have *: "\<And>x. f l = f x \<longleftrightarrow> f x = f l" by auto | |
| 68386 | 3062 | 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 | 3063 | unfolding filter_sort by (rule properties_for_sort_key) (auto intro: sorted_map_same) | 
| 68386 | 3064 | with * have **: "[x \<leftarrow> sort_key f xs . f l = f x] = [x \<leftarrow> xs. f l = f x]" by simp | 
| 40346 | 3065 | have "\<And>x P. P (f x) ?pivot \<and> f l = f x \<longleftrightarrow> P (f l) ?pivot \<and> f l = f x" by auto | 
| 68386 | 3066 | then have "\<And>P. [x \<leftarrow> sort_key f xs . P (f x) ?pivot \<and> f l = f x] = | 
| 3067 | [x \<leftarrow> sort_key f xs. P (f l) ?pivot \<and> f l = f x]" by simp | |
| 67398 | 3068 | note *** = this [of "(<)"] this [of "(>)"] this [of "(=)"] | 
| 68386 | 3069 | 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 | 3070 | proof (cases "f l" ?pivot rule: linorder_cases) | 
| 46730 | 3071 | case less | 
| 3072 | then have "f l \<noteq> ?pivot" and "\<not> f l > ?pivot" by auto | |
| 3073 | with less show ?thesis | |
| 40346 | 3074 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 3075 | next | 
| 40306 | 3076 | case equal then show ?thesis | 
| 40346 | 3077 | by (simp add: * less_le) | 
| 40305 
41833242cc42
tuned lemma proposition of properties_for_sort_key
 haftmann parents: 
40303diff
changeset | 3078 | next | 
| 46730 | 3079 | case greater | 
| 3080 | then have "f l \<noteq> ?pivot" and "\<not> f l < ?pivot" by auto | |
| 3081 | with greater show ?thesis | |
| 40346 | 3082 | by (simp add: filter_sort [symmetric] ** ***) | 
| 40306 | 3083 | qed | 
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3084 | qed | 
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3085 | |
| 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3086 | lemma sort_by_quicksort: | 
| 68386 | 3087 | "sort xs = sort [x\<leftarrow>xs. x < xs ! (length xs div 2)] | 
| 3088 | @ [x\<leftarrow>xs. x = xs ! (length xs div 2)] | |
| 3089 | @ sort [x\<leftarrow>xs. x > xs ! (length xs div 2)]" (is "sort ?lhs = ?rhs") | |
| 40303 
2d507370e879
lemmas multiset_of_filter, sort_key_by_quicksort
 haftmann parents: 
40250diff
changeset | 3090 | 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 | 3091 | |
| 68983 | 3092 | text \<open>A stable parameterized quicksort\<close> | 
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3093 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3094 | definition part :: "('b \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'b list \<Rightarrow> 'b list \<times> 'b list \<times> 'b list" where
 | 
| 68386 | 3095 | "part f pivot xs = ([x \<leftarrow> xs. f x < pivot], [x \<leftarrow> xs. f x = pivot], [x \<leftarrow> xs. pivot < f x])" | 
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3096 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3097 | lemma part_code [code]: | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3098 | "part f pivot [] = ([], [], [])" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3099 | "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 | 3100 | if x' < pivot then (x # lts, eqs, gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3101 | else if x' > pivot then (lts, eqs, x # gts) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3102 | else (lts, x # eqs, gts))" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3103 | by (auto simp add: part_def Let_def split_def) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3104 | |
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3105 | lemma sort_key_by_quicksort_code [code]: | 
| 60606 | 3106 | "sort_key f xs = | 
| 3107 | (case xs of | |
| 3108 | [] \<Rightarrow> [] | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3109 | | [x] \<Rightarrow> xs | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3110 | | [x, y] \<Rightarrow> (if f x \<le> f y then xs else [y, x]) | 
| 60606 | 3111 | | _ \<Rightarrow> | 
| 3112 | let (lts, eqs, gts) = part f (f (xs ! (length xs div 2))) xs | |
| 3113 | in sort_key f lts @ eqs @ sort_key f gts)" | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3114 | proof (cases xs) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3115 | case Nil then show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3116 | next | 
| 46921 | 3117 | case (Cons _ ys) note hyps = Cons show ?thesis | 
| 3118 | proof (cases ys) | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3119 | case Nil with hyps show ?thesis by simp | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3120 | next | 
| 46921 | 3121 | case (Cons _ zs) note hyps = hyps Cons show ?thesis | 
| 3122 | proof (cases zs) | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3123 | case Nil with hyps show ?thesis by auto | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3124 | next | 
| 58425 | 3125 | case Cons | 
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3126 | from sort_key_by_quicksort [of f xs] | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3127 | 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 | 3128 | in sort_key f lts @ eqs @ sort_key f gts)" | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3129 | 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 | 3130 | with hyps Cons show ?thesis by (simp only: list.cases) | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3131 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3132 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3133 | qed | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3134 | |
| 39533 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 3135 | end | 
| 
91a0ff0ff237
generalized lemmas multiset_of_insort, multiset_of_sort, properties_for_sort for *_key variants
 haftmann parents: 
39314diff
changeset | 3136 | |
| 40347 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3137 | hide_const (open) part | 
| 
429bf4388b2f
added code lemmas for stable parametrized quicksort
 haftmann parents: 
40346diff
changeset | 3138 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 3139 | lemma mset_remdups_subset_eq: "mset (remdups xs) \<subseteq># mset xs" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 3140 | 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 | 3141 | |
| 60515 | 3142 | lemma mset_update: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3143 |   "i < length ls \<Longrightarrow> mset (ls[i := v]) = add_mset v (mset ls - {#ls ! i#})"
 | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 3144 | 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 | 3145 | 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 | 3146 | 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 | 3147 | 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 | 3148 | 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 | 3149 | 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 | 3150 | 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 | 3151 | 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 | 3152 | 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 | 3153 | with Cons show ?thesis | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3154 | by (cases \<open>x = xs ! i'\<close>) auto | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 3155 | 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 | 3156 | 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 | 3157 | |
| 60515 | 3158 | 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 | 3159 | "i < length ls \<Longrightarrow> j < length ls \<Longrightarrow> | 
| 60515 | 3160 | mset (ls[j := ls ! i, i := ls ! j]) = mset ls" | 
| 3161 | 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 | 3162 | |
| 73327 
fd32f08f4fb5
more connections between mset _ = mset _ and permutations
 haftmann parents: 
73301diff
changeset | 3163 | lemma mset_eq_finite: | 
| 73301 | 3164 |   \<open>finite {ys. mset ys = mset xs}\<close>
 | 
| 3165 | proof - | |
| 3166 |   have \<open>{ys. mset ys = mset xs} \<subseteq> {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close>
 | |
| 3167 | by (auto simp add: dest: mset_eq_setD mset_eq_length) | |
| 3168 |   moreover have \<open>finite {ys. set ys \<subseteq> set xs \<and> length ys \<le> length xs}\<close>
 | |
| 3169 | using finite_lists_length_le by blast | |
| 3170 | ultimately show ?thesis | |
| 3171 | by (rule finite_subset) | |
| 3172 | qed | |
| 3173 | ||
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 3174 | |
| 60500 | 3175 | subsection \<open>The multiset order\<close> | 
| 3176 | ||
| 60606 | 3177 | definition mult1 :: "('a \<times> 'a) set \<Rightarrow> ('a multiset \<times> 'a multiset) set" where
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3178 |   "mult1 r = {(N, M). \<exists>a M0 K. M = add_mset a M0 \<and> N = M0 + K \<and>
 | 
| 60607 | 3179 | (\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r)}" | 
| 60606 | 3180 | |
| 3181 | definition mult :: "('a \<times> 'a) set \<Rightarrow> ('a multiset \<times> 'a multiset) set" where
 | |
| 37765 | 3182 | "mult r = (mult1 r)\<^sup>+" | 
| 10249 | 3183 | |
| 74858 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 desharna parents: 
74806diff
changeset | 3184 | definition multp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
 | 
| 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 desharna parents: 
74806diff
changeset | 3185 |   "multp r M N \<longleftrightarrow> (M, N) \<in> mult {(x, y). r x y}"
 | 
| 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 desharna parents: 
74806diff
changeset | 3186 | |
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3187 | declare multp_def[pred_set_conv] | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3188 | |
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3189 | lemma mult1I: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3190 | assumes "M = add_mset a M0" and "N = M0 + K" and "\<And>b. b \<in># K \<Longrightarrow> (b, a) \<in> r" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3191 | shows "(N, M) \<in> mult1 r" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3192 | using assms unfolding mult1_def by blast | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3193 | |
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3194 | lemma mult1E: | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3195 | assumes "(N, M) \<in> mult1 r" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3196 | obtains a M0 K where "M = add_mset a M0" "N = M0 + K" "\<And>b. b \<in># K \<Longrightarrow> (b, a) \<in> r" | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3197 | using assms unfolding mult1_def by blast | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 3198 | |
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3199 | lemma mono_mult1: | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3200 | assumes "r \<subseteq> r'" shows "mult1 r \<subseteq> mult1 r'" | 
| 74858 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 desharna parents: 
74806diff
changeset | 3201 | unfolding mult1_def using assms by blast | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3202 | |
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3203 | lemma mono_mult: | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3204 | assumes "r \<subseteq> r'" shows "mult r \<subseteq> mult r'" | 
| 74858 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 desharna parents: 
74806diff
changeset | 3205 | unfolding mult_def using mono_mult1[OF assms] trancl_mono by blast | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3206 | |
| 74859 | 3207 | lemma mono_multp[mono]: "r \<le> r' \<Longrightarrow> multp r \<le> multp r'" | 
| 3208 | unfolding le_fun_def le_bool_def | |
| 3209 | proof (intro allI impI) | |
| 3210 | fix M N :: "'a multiset" | |
| 3211 | assume "\<forall>x xa. r x xa \<longrightarrow> r' x xa" | |
| 3212 |   hence "{(x, y). r x y} \<subseteq> {(x, y). r' x y}"
 | |
| 3213 | by blast | |
| 3214 | thus "multp r M N \<Longrightarrow> multp r' M N" | |
| 3215 | unfolding multp_def | |
| 3216 | by (fact mono_mult[THEN subsetD, rotated]) | |
| 3217 | qed | |
| 3218 | ||
| 23751 | 3219 | lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
 | 
| 74858 
c5059f9fbfba
added Multiset.multp as predicate equivalent of Multiset.mult
 desharna parents: 
74806diff
changeset | 3220 | by (simp add: mult1_def) | 
| 10249 | 3221 | |
| 74860 | 3222 | |
| 3223 | subsubsection \<open>Well-foundedness\<close> | |
| 3224 | ||
| 60608 | 3225 | lemma less_add: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3226 | assumes mult1: "(N, add_mset a M0) \<in> mult1 r" | 
| 60608 | 3227 | shows | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3228 | "(\<exists>M. (M, M0) \<in> mult1 r \<and> N = add_mset a M) \<or> | 
| 60608 | 3229 | (\<exists>K. (\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r) \<and> N = M0 + K)" | 
| 3230 | proof - | |
| 60607 | 3231 | let ?r = "\<lambda>K a. \<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3232 | let ?R = "\<lambda>N M. \<exists>a M0 K. M = add_mset a M0 \<and> N = M0 + K \<and> ?r K a" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3233 | obtain a' M0' K where M0: "add_mset a M0 = add_mset a' M0'" | 
| 60608 | 3234 | and N: "N = M0' + K" | 
| 3235 | and r: "?r K a'" | |
| 3236 | using mult1 unfolding mult1_def by auto | |
| 3237 | show ?thesis (is "?case1 \<or> ?case2") | |
| 60606 | 3238 | proof - | 
| 3239 | from M0 consider "M0 = M0'" "a = a'" | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3240 | | K' where "M0 = add_mset a' K'" "M0' = add_mset a K'" | 
| 60606 | 3241 | by atomize_elim (simp only: add_eq_conv_ex) | 
| 18258 | 3242 | then show ?thesis | 
| 60606 | 3243 | proof cases | 
| 3244 | case 1 | |
| 11464 | 3245 | with N r have "?r K a \<and> N = M0 + K" by simp | 
| 60606 | 3246 | then have ?case2 .. | 
| 3247 | then show ?thesis .. | |
| 10249 | 3248 | next | 
| 60606 | 3249 | case 2 | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 3250 | from N 2(2) have n: "N = add_mset a (K' + K)" by simp | 
| 60606 | 3251 | with r 2(1) have "?R (K' + K) M0" by blast | 
| 60608 | 3252 | with n have ?case1 by (simp add: mult1_def) | 
| 60606 | 3253 | then show ?thesis .. | 
| 10249 | 3254 | qed | 
| 3255 | qed | |
| 3256 | qed | |
| 3257 | ||
| 60608 | 3258 | lemma all_accessible: | 
| 3259 | assumes "wf r" | |
| 3260 | shows "\<forall>M. M \<in> Wellfounded.acc (mult1 r)" | |
| 10249 | 3261 | proof | 
| 3262 | let ?R = "mult1 r" | |
| 54295 | 3263 | let ?W = "Wellfounded.acc ?R" | 
| 10249 | 3264 |   {
 | 
| 3265 | fix M M0 a | |
| 23751 | 3266 | assume M0: "M0 \<in> ?W" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3267 | and wf_hyp: "\<And>b. (b, a) \<in> r \<Longrightarrow> (\<forall>M \<in> ?W. add_mset b M \<in> ?W)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3268 | and acc_hyp: "\<forall>M. (M, M0) \<in> ?R \<longrightarrow> add_mset a M \<in> ?W" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3269 | have "add_mset a M0 \<in> ?W" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3270 | proof (rule accI [of "add_mset a M0"]) | 
| 10249 | 3271 | fix N | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3272 | assume "(N, add_mset a M0) \<in> ?R" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3273 | then consider M where "(M, M0) \<in> ?R" "N = add_mset a M" | 
| 60608 | 3274 | | K where "\<forall>b. b \<in># K \<longrightarrow> (b, a) \<in> r" "N = M0 + K" | 
| 3275 | by atomize_elim (rule less_add) | |
| 23751 | 3276 | then show "N \<in> ?W" | 
| 60608 | 3277 | proof cases | 
| 3278 | case 1 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3279 | from acc_hyp have "(M, M0) \<in> ?R \<longrightarrow> add_mset a M \<in> ?W" .. | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3280 | from this and \<open>(M, M0) \<in> ?R\<close> have "add_mset a M \<in> ?W" .. | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3281 | then show "N \<in> ?W" by (simp only: \<open>N = add_mset a M\<close>) | 
| 10249 | 3282 | next | 
| 60608 | 3283 | case 2 | 
| 3284 | from this(1) have "M0 + K \<in> ?W" | |
| 10249 | 3285 | proof (induct K) | 
| 18730 | 3286 | case empty | 
| 23751 | 3287 |           from M0 show "M0 + {#} \<in> ?W" by simp
 | 
| 18730 | 3288 | next | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3289 | case (add x K) | 
| 23751 | 3290 | from add.prems have "(x, a) \<in> r" by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3291 | with wf_hyp have "\<forall>M \<in> ?W. add_mset x M \<in> ?W" by blast | 
| 23751 | 3292 | moreover from add have "M0 + K \<in> ?W" by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3293 | ultimately have "add_mset x (M0 + K) \<in> ?W" .. | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3294 | then show "M0 + (add_mset x K) \<in> ?W" by simp | 
| 10249 | 3295 | qed | 
| 60608 | 3296 | then show "N \<in> ?W" by (simp only: 2(2)) | 
| 10249 | 3297 | qed | 
| 3298 | qed | |
| 3299 | } note tedious_reasoning = this | |
| 3300 | ||
| 60608 | 3301 | show "M \<in> ?W" for M | 
| 10249 | 3302 | proof (induct M) | 
| 23751 | 3303 |     show "{#} \<in> ?W"
 | 
| 10249 | 3304 | proof (rule accI) | 
| 23751 | 3305 |       fix b assume "(b, {#}) \<in> ?R"
 | 
| 3306 | with not_less_empty show "b \<in> ?W" by contradiction | |
| 10249 | 3307 | qed | 
| 3308 | ||
| 23751 | 3309 | fix M a assume "M \<in> ?W" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3310 | from \<open>wf r\<close> have "\<forall>M \<in> ?W. add_mset a M \<in> ?W" | 
| 10249 | 3311 | proof induct | 
| 3312 | fix a | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3313 | assume r: "\<And>b. (b, a) \<in> r \<Longrightarrow> (\<forall>M \<in> ?W. add_mset b M \<in> ?W)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3314 | show "\<forall>M \<in> ?W. add_mset a M \<in> ?W" | 
| 10249 | 3315 | proof | 
| 23751 | 3316 | fix M assume "M \<in> ?W" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3317 | then show "add_mset a M \<in> ?W" | 
| 23373 | 3318 | by (rule acc_induct) (rule tedious_reasoning [OF _ r]) | 
| 10249 | 3319 | qed | 
| 3320 | qed | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3321 | from this and \<open>M \<in> ?W\<close> show "add_mset a M \<in> ?W" .. | 
| 10249 | 3322 | qed | 
| 3323 | qed | |
| 3324 | ||
| 74860 | 3325 | lemma wf_mult1: "wf r \<Longrightarrow> wf (mult1 r)" | 
| 3326 | by (rule acc_wfI) (rule all_accessible) | |
| 3327 | ||
| 3328 | lemma wf_mult: "wf r \<Longrightarrow> wf (mult r)" | |
| 3329 | unfolding mult_def by (rule wf_trancl) (rule wf_mult1) | |
| 3330 | ||
| 80324 | 3331 | lemma wfp_multp: "wfp r \<Longrightarrow> wfp (multp r)" | 
| 80322 | 3332 | unfolding multp_def wfp_def | 
| 74860 | 3333 | by (simp add: wf_mult) | 
| 10249 | 3334 | |
| 3335 | ||
| 60500 | 3336 | subsubsection \<open>Closure-free presentation\<close> | 
| 3337 | ||
| 3338 | text \<open>One direction.\<close> | |
| 10249 | 3339 | lemma mult_implies_one_step: | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3340 | assumes | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3341 | trans: "trans r" and | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3342 | MN: "(M, N) \<in> mult r" | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3343 |   shows "\<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and> (\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> r)"
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3344 | using MN unfolding mult_def mult1_def | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3345 | proof (induction rule: converse_trancl_induct) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3346 | case (base y) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3347 | then show ?case by force | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3348 | next | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3349 | case (step y z) note yz = this(1) and zN = this(2) and N_decomp = this(3) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3350 | obtain I J K where | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3351 |     N: "N = I + J" "z = I + K" "J \<noteq> {#}" "\<forall>k\<in>#K. \<exists>j\<in>#J. (k, j) \<in> r"
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3352 | using N_decomp by blast | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3353 | obtain a M0 K' where | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3354 | z: "z = add_mset a M0" and y: "y = M0 + K'" and K: "\<forall>b. b \<in># K' \<longrightarrow> (b, a) \<in> r" | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3355 | using yz by blast | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3356 | show ?case | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3357 | proof (cases "a \<in># K") | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3358 | case True | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3359 | moreover have "\<exists>j\<in>#J. (k, j) \<in> r" if "k \<in># K'" for k | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3360 | using K N trans True by (meson that transE) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3361 | ultimately show ?thesis | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3362 |       by (rule_tac x = I in exI, rule_tac x = J in exI, rule_tac x = "(K - {#a#}) + K'" in exI)
 | 
| 64017 
6e7bf7678518
more multiset simp rules
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63924diff
changeset | 3363 | (use z y N in \<open>auto simp del: subset_mset.add_diff_assoc2 dest: in_diffD\<close>) | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3364 | next | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3365 | case False | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3366 | then have "a \<in># I" by (metis N(2) union_iff union_single_eq_member z) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3367 |     moreover have "M0 = I + K - {#a#}"
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3368 | using N(2) z by force | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3369 | ultimately show ?thesis | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3370 |       by (rule_tac x = "I - {#a#}" in exI, rule_tac x = "add_mset a J" in exI,
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3371 | rule_tac x = "K + K'" in exI) | 
| 64017 
6e7bf7678518
more multiset simp rules
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63924diff
changeset | 3372 | (use z y N False K in \<open>auto simp: add.assoc\<close>) | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3373 | qed | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3374 | qed | 
| 10249 | 3375 | |
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3376 | lemma multp_implies_one_step: | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3377 |   "transp R \<Longrightarrow> multp R M N \<Longrightarrow> \<exists>I J K. N = I + J \<and> M = I + K \<and> J \<noteq> {#} \<and> (\<forall>k\<in>#K. \<exists>x\<in>#J. R k x)"
 | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3378 | by (rule mult_implies_one_step[to_pred]) | 
| 74861 
74ac414618e2
added lemmas multp_implies_one_step, one_step_implies_multp, and subset_implies_multp
 desharna parents: 
74860diff
changeset | 3379 | |
| 17161 | 3380 | lemma one_step_implies_mult: | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3381 | assumes | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3382 |     "J \<noteq> {#}" and
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3383 | "\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> r" | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3384 | shows "(I + K, I + J) \<in> mult r" | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3385 | using assms | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3386 | proof (induction "size J" arbitrary: I J K) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3387 | case 0 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3388 | then show ?case by auto | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3389 | next | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3390 | case (Suc n) note IH = this(1) and size_J = this(2)[THEN sym] | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3391 | obtain J' a where J: "J = add_mset a J'" | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3392 | using size_J by (blast dest: size_eq_Suc_imp_eq_union) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3393 | show ?case | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3394 |   proof (cases "J' = {#}")
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3395 | case True | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3396 | then show ?thesis | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3397 | using J Suc by (fastforce simp add: mult_def mult1_def) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3398 | next | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3399 | case [simp]: False | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3400 |     have K: "K = {#x \<in># K. (x, a) \<in> r#} + {#x \<in># K. (x, a) \<notin> r#}"
 | 
| 68992 | 3401 | by simp | 
| 63795 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3402 |     have "(I + K, (I + {# x \<in># K. (x, a) \<in> r #}) + J') \<in> mult r"
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3403 |       using IH[of J' "{# x \<in># K. (x, a) \<notin> r#}" "I + {# x \<in># K. (x, a) \<in> r#}"]
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3404 | J Suc.prems K size_J by (auto simp: ac_simps) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3405 |     moreover have "(I + {#x \<in># K. (x, a) \<in> r#} + J', I + J) \<in> mult r"
 | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3406 | by (fastforce simp: J mult1_def mult_def) | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3407 | ultimately show ?thesis | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3408 | unfolding mult_def by simp | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3409 | qed | 
| 
7f6128adfe67
tuning multisets; more interpretations
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63794diff
changeset | 3410 | qed | 
| 10249 | 3411 | |
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3412 | lemma one_step_implies_multp: | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3413 |   "J \<noteq> {#} \<Longrightarrow> \<forall>k\<in>#K. \<exists>j\<in>#J. R k j \<Longrightarrow> multp R (I + K) (I + J)"
 | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3414 |   by (rule one_step_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def, simplified])
 | 
| 74861 
74ac414618e2
added lemmas multp_implies_one_step, one_step_implies_multp, and subset_implies_multp
 desharna parents: 
74860diff
changeset | 3415 | |
| 65047 | 3416 | lemma subset_implies_mult: | 
| 3417 | assumes sub: "A \<subset># B" | |
| 3418 | shows "(A, B) \<in> mult r" | |
| 3419 | proof - | |
| 3420 | have ApBmA: "A + (B - A) = B" | |
| 3421 | using sub by simp | |
| 3422 |   have BmA: "B - A \<noteq> {#}"
 | |
| 3423 | using sub by (simp add: Diff_eq_empty_iff_mset subset_mset.less_le_not_le) | |
| 3424 | thus ?thesis | |
| 3425 |     by (rule one_step_implies_mult[of "B - A" "{#}" _ A, unfolded ApBmA, simplified])
 | |
| 3426 | qed | |
| 3427 | ||
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3428 | lemma subset_implies_multp: "A \<subset># B \<Longrightarrow> multp r A B" | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3429 |   by (rule subset_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def])
 | 
| 74861 
74ac414618e2
added lemmas multp_implies_one_step, one_step_implies_multp, and subset_implies_multp
 desharna parents: 
74860diff
changeset | 3430 | |
| 77688 | 3431 | lemma multp_repeat_mset_repeat_msetI: | 
| 3432 | assumes "transp R" and "multp R A B" and "n \<noteq> 0" | |
| 3433 | shows "multp R (repeat_mset n A) (repeat_mset n B)" | |
| 3434 | proof - | |
| 3435 | from \<open>transp R\<close> \<open>multp R A B\<close> obtain I J K where | |
| 3436 |     "B = I + J" and "A = I + K" and "J \<noteq> {#}" and "\<forall>k \<in># K. \<exists>x \<in># J. R k x"
 | |
| 3437 | by (auto dest: multp_implies_one_step) | |
| 3438 | ||
| 3439 | have repeat_n_A_eq: "repeat_mset n A = repeat_mset n I + repeat_mset n K" | |
| 3440 | using \<open>A = I + K\<close> by simp | |
| 3441 | ||
| 3442 | have repeat_n_B_eq: "repeat_mset n B = repeat_mset n I + repeat_mset n J" | |
| 3443 | using \<open>B = I + J\<close> by simp | |
| 3444 | ||
| 3445 | show ?thesis | |
| 3446 | unfolding repeat_n_A_eq repeat_n_B_eq | |
| 3447 | proof (rule one_step_implies_multp) | |
| 3448 |     from \<open>n \<noteq> 0\<close> show "repeat_mset n J \<noteq> {#}"
 | |
| 3449 |       using \<open>J \<noteq> {#}\<close>
 | |
| 3450 | by (simp add: repeat_mset_eq_empty_iff) | |
| 3451 | next | |
| 3452 | show "\<forall>k \<in># repeat_mset n K. \<exists>j \<in># repeat_mset n J. R k j" | |
| 3453 | using \<open>\<forall>k \<in># K. \<exists>x \<in># J. R k x\<close> | |
| 3454 | by (metis count_greater_zero_iff nat_0_less_mult_iff repeat_mset.rep_eq) | |
| 3455 | qed | |
| 3456 | qed | |
| 3457 | ||
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3458 | |
| 75560 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3459 | subsubsection \<open>Monotonicity\<close> | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3460 | |
| 76401 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3461 | lemma multp_mono_strong: | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3462 | assumes "multp R M1 M2" and "transp R" and | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3463 | S_if_R: "\<And>x y. x \<in> set_mset M1 \<Longrightarrow> y \<in> set_mset M2 \<Longrightarrow> R x y \<Longrightarrow> S x y" | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3464 | shows "multp S M1 M2" | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3465 | proof - | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3466 |   obtain I J K where "M2 = I + J" and "M1 = I + K" and "J \<noteq> {#}" and "\<forall>k\<in>#K. \<exists>x\<in>#J. R k x"
 | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3467 | using multp_implies_one_step[OF \<open>transp R\<close> \<open>multp R M1 M2\<close>] by auto | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3468 | show ?thesis | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3469 | unfolding \<open>M2 = I + J\<close> \<open>M1 = I + K\<close> | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3470 |   proof (rule one_step_implies_multp[OF \<open>J \<noteq> {#}\<close>])
 | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3471 | show "\<forall>k\<in>#K. \<exists>j\<in>#J. S k j" | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3472 | using S_if_R | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3473 | by (metis \<open>M1 = I + K\<close> \<open>M2 = I + J\<close> \<open>\<forall>k\<in>#K. \<exists>x\<in>#J. R k x\<close> union_iff) | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3474 | qed | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3475 | qed | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3476 | |
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3477 | lemma mult_mono_strong: | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3478 | assumes "(M1, M2) \<in> mult r" and "trans r" and | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3479 | S_if_R: "\<And>x y. x \<in> set_mset M1 \<Longrightarrow> y \<in> set_mset M2 \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s" | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3480 | shows "(M1, M2) \<in> mult s" | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3481 | using assms multp_mono_strong[of "\<lambda>x y. (x, y) \<in> r" M1 M2 "\<lambda>x y. (x, y) \<in> s", | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3482 | unfolded multp_def transp_trans_eq, simplified] | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3483 | by blast | 
| 
e7e8fbc89870
added lemmas multp_mono_strong and mult_mono_strong
 desharna parents: 
76359diff
changeset | 3484 | |
| 75584 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3485 | lemma monotone_on_multp_multp_image_mset: | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3486 | assumes "monotone_on A orda ordb f" and "transp orda" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3487 |   shows "monotone_on {M. set_mset M \<subseteq> A} (multp orda) (multp ordb) (image_mset f)"
 | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3488 | proof (rule monotone_onI) | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3489 | fix M1 M2 | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3490 | assume | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3491 |     M1_in: "M1 \<in> {M. set_mset M \<subseteq> A}" and
 | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3492 |     M2_in: "M2 \<in> {M. set_mset M \<subseteq> A}" and
 | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3493 | M1_lt_M2: "multp orda M1 M2" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3494 | |
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3495 | from multp_implies_one_step[OF \<open>transp orda\<close> M1_lt_M2] obtain I J K where | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3496 | M2_eq: "M2 = I + J" and | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3497 | M1_eq: "M1 = I + K" and | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3498 |     J_neq_mempty: "J \<noteq> {#}" and
 | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3499 | ball_K_less: "\<forall>k\<in>#K. \<exists>x\<in>#J. orda k x" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3500 | by metis | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3501 | |
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3502 | have "multp ordb (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3503 | proof (intro one_step_implies_multp ballI) | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3504 |     show "image_mset f J \<noteq> {#}"
 | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3505 | using J_neq_mempty by simp | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3506 | next | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3507 | fix k' assume "k'\<in>#image_mset f K" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3508 | then obtain k where "k' = f k" and k_in: "k \<in># K" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3509 | by auto | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3510 | then obtain j where j_in: "j\<in>#J" and "orda k j" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3511 | using ball_K_less by auto | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3512 | |
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3513 | have "ordb (f k) (f j)" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3514 | proof (rule \<open>monotone_on A orda ordb f\<close>[THEN monotone_onD, OF _ _ \<open>orda k j\<close>]) | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3515 | show "k \<in> A" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3516 | using M1_eq M1_in k_in by auto | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3517 | next | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3518 | show "j \<in> A" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3519 | using M2_eq M2_in j_in by auto | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3520 | qed | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3521 | thus "\<exists>j\<in>#image_mset f J. ordb k' j" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3522 | using \<open>j \<in># J\<close> \<open>k' = f k\<close> by auto | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3523 | qed | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3524 | thus "multp ordb (image_mset f M1) (image_mset f M2)" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3525 | by (simp add: M1_eq M2_eq) | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3526 | qed | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3527 | |
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3528 | lemma monotone_multp_multp_image_mset: | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3529 | assumes "monotone orda ordb f" and "transp orda" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3530 | shows "monotone (multp orda) (multp ordb) (image_mset f)" | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3531 | by (rule monotone_on_multp_multp_image_mset[OF assms, simplified]) | 
| 
c32658b9e4df
added lemmas monotone{,_on}_multp_multp_image_mset
 desharna parents: 
75560diff
changeset | 3532 | |
| 77832 | 3533 | lemma multp_image_mset_image_msetI: | 
| 3534 | assumes "multp (\<lambda>x y. R (f x) (f y)) M1 M2" and "transp R" | |
| 3535 | shows "multp R (image_mset f M1) (image_mset f M2)" | |
| 3536 | proof - | |
| 3537 | from \<open>transp R\<close> have "transp (\<lambda>x y. R (f x) (f y))" | |
| 3538 | by (auto intro: transpI dest: transpD) | |
| 3539 | with \<open>multp (\<lambda>x y. R (f x) (f y)) M1 M2\<close> obtain I J K where | |
| 3540 |     "M2 = I + J" and "M1 = I + K" and "J \<noteq> {#}" and "\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)"
 | |
| 3541 | using multp_implies_one_step by blast | |
| 3542 | ||
| 3543 | have "multp R (image_mset f I + image_mset f K) (image_mset f I + image_mset f J)" | |
| 3544 | proof (rule one_step_implies_multp) | |
| 3545 |     show "image_mset f J \<noteq> {#}"
 | |
| 3546 |       by (simp add: \<open>J \<noteq> {#}\<close>)
 | |
| 3547 | next | |
| 3548 | show "\<forall>k\<in>#image_mset f K. \<exists>j\<in>#image_mset f J. R k j" | |
| 3549 | by (simp add: \<open>\<forall>k\<in>#K. \<exists>x\<in>#J. R (f k) (f x)\<close>) | |
| 3550 | qed | |
| 3551 | thus ?thesis | |
| 3552 | by (simp add: \<open>M1 = I + K\<close> \<open>M2 = I + J\<close>) | |
| 3553 | qed | |
| 3554 | ||
| 75560 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3555 | lemma multp_image_mset_image_msetD: | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3556 | assumes | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3557 | "multp R (image_mset f A) (image_mset f B)" and | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3558 | "transp R" and | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3559 | inj_on_f: "inj_on f (set_mset A \<union> set_mset B)" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3560 | shows "multp (\<lambda>x y. R (f x) (f y)) A B" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3561 | proof - | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3562 | from assms(1,2) obtain I J K where | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3563 | f_B_eq: "image_mset f B = I + J" and | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3564 | f_A_eq: "image_mset f A = I + K" and | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3565 |     J_neq_mempty: "J \<noteq> {#}" and
 | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3566 | ball_K_less: "\<forall>k\<in>#K. \<exists>x\<in>#J. R k x" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3567 | by (auto dest: multp_implies_one_step) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3568 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3569 | from f_B_eq obtain I' J' where | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3570 | B_def: "B = I' + J'" and I_def: "I = image_mset f I'" and J_def: "J = image_mset f J'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3571 | using image_mset_eq_plusD by blast | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3572 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3573 | from inj_on_f have inj_on_f': "inj_on f (set_mset A \<union> set_mset I')" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3574 | by (rule inj_on_subset) (auto simp add: B_def) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3575 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3576 | from f_A_eq obtain K' where | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3577 | A_def: "A = I' + K'" and K_def: "K = image_mset f K'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3578 | by (auto simp: I_def dest: image_mset_eq_image_mset_plusD[OF _ inj_on_f']) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3579 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3580 | show ?thesis | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3581 | unfolding A_def B_def | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3582 | proof (intro one_step_implies_multp ballI) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3583 |     from J_neq_mempty show "J' \<noteq> {#}"
 | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3584 | by (simp add: J_def) | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3585 | next | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3586 | fix k assume "k \<in># K'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3587 | with ball_K_less obtain j' where "j' \<in># J" and "R (f k) j'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3588 | using K_def by auto | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3589 | moreover then obtain j where "j \<in># J'" and "f j = j'" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3590 | using J_def by auto | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3591 | ultimately show "\<exists>j\<in>#J'. R (f k) (f j)" | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3592 | by blast | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3593 | qed | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3594 | qed | 
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3595 | |
| 
aeb797356de0
added lemmas image_mset_eq_{image_mset_plus,plus,plus_image_mset}D, and multp_image_mset_image_msetD
 desharna parents: 
75467diff
changeset | 3596 | |
| 74862 
aa51e974b688
added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
 desharna parents: 
74861diff
changeset | 3597 | subsubsection \<open>The multiset extension is cancellative for multiset union\<close> | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3598 | |
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3599 | lemma mult_cancel: | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3600 | assumes "trans s" and "irrefl_on (set_mset Z) s" | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3601 | shows "(X + Z, Y + Z) \<in> mult s \<longleftrightarrow> (X, Y) \<in> mult s" (is "?L \<longleftrightarrow> ?R") | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3602 | proof | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3603 | assume ?L thus ?R | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3604 | using \<open>irrefl_on (set_mset Z) s\<close> | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3605 | proof (induct Z) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3606 | case (add z Z) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3607 |     obtain X' Y' Z' where *: "add_mset z X + Z = Z' + X'" "add_mset z Y + Z = Z' + Y'" "Y' \<noteq> {#}"
 | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3608 | "\<forall>x \<in> set_mset X'. \<exists>y \<in> set_mset Y'. (x, y) \<in> s" | 
| 64911 | 3609 | using mult_implies_one_step[OF \<open>trans s\<close> add(2)] by auto | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3610 | consider Z2 where "Z' = add_mset z Z2" | X2 Y2 where "X' = add_mset z X2" "Y' = add_mset z Y2" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 3611 | using *(1,2) by (metis add_mset_remove_trivial_If insert_iff set_mset_add_mset_insert union_iff) | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3612 | thus ?case | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3613 | proof (cases) | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3614 | case 1 thus ?thesis | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3615 | using * one_step_implies_mult[of Y' X' s Z2] add(3) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3616 |         by (auto simp: add.commute[of _ "{#_#}"] add.assoc intro: add(1) elim: irrefl_on_subset)
 | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3617 | next | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3618 | case 2 then obtain y where "y \<in> set_mset Y2" "(z, y) \<in> s" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3619 | using *(4) \<open>irrefl_on (set_mset (add_mset z Z)) s\<close> | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3620 | by (auto simp: irrefl_on_def) | 
| 64911 | 3621 | moreover from this transD[OF \<open>trans s\<close> _ this(2)] | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3622 | have "x' \<in> set_mset X2 \<Longrightarrow> \<exists>y \<in> set_mset Y2. (x', y) \<in> s" for x' | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3623 | using 2 *(4)[rule_format, of x'] by auto | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3624 | ultimately show ?thesis | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3625 | using * one_step_implies_mult[of Y2 X2 s Z'] 2 add(3) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3626 |         by (force simp: add.commute[of "{#_#}"] add.assoc[symmetric] intro: add(1)
 | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3627 | elim: irrefl_on_subset) | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3628 | qed | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3629 | qed auto | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3630 | next | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3631 | assume ?R then obtain I J K | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3632 |     where "Y = I + J" "X = I + K" "J \<noteq> {#}" "\<forall>k \<in> set_mset K. \<exists>j \<in> set_mset J. (k, j) \<in> s"
 | 
| 64911 | 3633 | using mult_implies_one_step[OF \<open>trans s\<close>] by blast | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3634 | thus ?L using one_step_implies_mult[of J K s "I + Z"] by (auto simp: ac_simps) | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3635 | qed | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3636 | |
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3637 | lemma multp_cancel: | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3638 | "transp R \<Longrightarrow> irreflp_on (set_mset Z) R \<Longrightarrow> multp R (X + Z) (Y + Z) \<longleftrightarrow> multp R X Y" | 
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3639 | by (rule mult_cancel[to_pred]) | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3640 | |
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3641 | lemma mult_cancel_add_mset: | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3642 |   "trans r \<Longrightarrow> irrefl_on {z} r \<Longrightarrow>
 | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3643 | ((add_mset z X, add_mset z Y) \<in> mult r) = ((X, Y) \<in> mult r)" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3644 |   by (rule mult_cancel[of _ "{#_#}", simplified])
 | 
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3645 | |
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3646 | lemma multp_cancel_add_mset: | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3647 |   "transp R \<Longrightarrow> irreflp_on {z} R \<Longrightarrow> multp R (add_mset z X) (add_mset z Y) = multp R X Y"
 | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3648 | by (rule mult_cancel_add_mset[to_pred, folded bot_set_def]) | 
| 74862 
aa51e974b688
added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
 desharna parents: 
74861diff
changeset | 3649 | |
| 74804 
5749fefd3fa0
simplified mult_cancel_max and introduced orginal lemma as mult_cancel_max0
 desharna parents: 
74803diff
changeset | 3650 | lemma mult_cancel_max0: | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3651 | assumes "trans s" and "irrefl_on (set_mset X \<inter> set_mset Y) s" | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3652 | shows "(X, Y) \<in> mult s \<longleftrightarrow> (X - X \<inter># Y, Y - X \<inter># Y) \<in> mult s" (is "?L \<longleftrightarrow> ?R") | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3653 | proof - | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3654 | have "(X - X \<inter># Y + X \<inter># Y, Y - X \<inter># Y + X \<inter># Y) \<in> mult s \<longleftrightarrow> (X - X \<inter># Y, Y - X \<inter># Y) \<in> mult s" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3655 | proof (rule mult_cancel) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3656 | from assms show "trans s" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3657 | by simp | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3658 | next | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3659 | from assms show "irrefl_on (set_mset (X \<inter># Y)) s" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3660 | by simp | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3661 | qed | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3662 | moreover have "X - X \<inter># Y + X \<inter># Y = X" "Y - X \<inter># Y + X \<inter># Y = Y" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3663 | by (auto simp flip: count_inject) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3664 | ultimately show ?thesis | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3665 | by simp | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3666 | qed | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3667 | |
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3668 | lemma mult_cancel_max: | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3669 | "trans r \<Longrightarrow> irrefl_on (set_mset X \<inter> set_mset Y) r \<Longrightarrow> | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3670 | (X, Y) \<in> mult r \<longleftrightarrow> (X - Y, Y - X) \<in> mult r" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3671 | by (rule mult_cancel_max0[simplified]) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3672 | |
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3673 | lemma multp_cancel_max: | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3674 | "transp R \<Longrightarrow> irreflp_on (set_mset X \<inter> set_mset Y) R \<Longrightarrow> multp R X Y \<longleftrightarrow> multp R (X - Y) (Y - X)" | 
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3675 | by (rule mult_cancel_max[to_pred]) | 
| 74862 
aa51e974b688
added lemmas multp_cancel, multp_cancel_add_mset, and multp_cancel_max
 desharna parents: 
74861diff
changeset | 3676 | |
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3677 | |
| 77049 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3678 | subsubsection \<open>Strict partial-order properties\<close> | 
| 74864 | 3679 | |
| 3680 | lemma mult1_lessE: | |
| 3681 |   assumes "(N, M) \<in> mult1 {(a, b). r a b}" and "asymp r"
 | |
| 3682 | obtains a M0 K where "M = add_mset a M0" "N = M0 + K" | |
| 3683 | "a \<notin># K" "\<And>b. b \<in># K \<Longrightarrow> r b a" | |
| 3684 | proof - | |
| 3685 | from assms obtain a M0 K where "M = add_mset a M0" "N = M0 + K" and | |
| 3686 | *: "b \<in># K \<Longrightarrow> r b a" for b by (blast elim: mult1E) | |
| 3687 | moreover from * [of a] have "a \<notin># K" | |
| 76682 
e260dabc88e6
added predicates asym_on and asymp_on and redefined asym and asymp to be abbreviations
 desharna parents: 
76611diff
changeset | 3688 | using \<open>asymp r\<close> by (meson asympD) | 
| 74864 | 3689 | ultimately show thesis by (auto intro: that) | 
| 3690 | qed | |
| 3691 | ||
| 79575 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3692 | lemma trans_on_mult: | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3693 | assumes "trans_on A r" and "\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A" | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3694 | shows "trans_on B (mult r)" | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3695 | using assms by (metis mult_def subset_UNIV trans_on_subset trans_trancl) | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3696 | |
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3697 | lemma trans_mult: "trans r \<Longrightarrow> trans (mult r)" | 
| 79575 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3698 | using trans_on_mult[of UNIV r UNIV, simplified] . | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3699 | |
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3700 | lemma transp_on_multp: | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3701 | assumes "transp_on A r" and "\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A" | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3702 | shows "transp_on B (multp r)" | 
| 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3703 | by (metis mult_def multp_def transD trans_trancl transp_onI) | 
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3704 | |
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3705 | lemma transp_multp: "transp r \<Longrightarrow> transp (multp r)" | 
| 79575 
b21d8401f0ca
added lemmas Multiset.transp_on_multp and Multiset.trans_on_mult
 desharna parents: 
78099diff
changeset | 3706 | using transp_on_multp[of UNIV r UNIV, simplified] . | 
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3707 | |
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3708 | lemma irrefl_mult: | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3709 | assumes "trans r" "irrefl r" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3710 | shows "irrefl (mult r)" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3711 | proof (intro irreflI notI) | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3712 | fix M | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3713 | assume "(M, M) \<in> mult r" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3714 | then obtain I J K where "M = I + J" and "M = I + K" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3715 |     and "J \<noteq> {#}" and "(\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset J. (k, j) \<in> r)"
 | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3716 | using mult_implies_one_step[OF \<open>trans r\<close>] by blast | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3717 |   then have *: "K \<noteq> {#}" and **: "\<forall>k\<in>set_mset K. \<exists>j\<in>set_mset K. (k, j) \<in> r" by auto
 | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3718 | have "finite (set_mset K)" by simp | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3719 |   hence "set_mset K = {}"
 | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3720 | using ** | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3721 | proof (induction rule: finite_induct) | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3722 | case empty | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3723 | thus ?case by simp | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3724 | next | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3725 | case (insert x F) | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3726 | have False | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3727 | using \<open>irrefl r\<close>[unfolded irrefl_def, rule_format] | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3728 | using \<open>trans r\<close>[THEN transD] | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3729 | by (metis equals0D insert.IH insert.prems insertE insertI1 insertI2) | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3730 | thus ?case .. | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3731 | qed | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3732 | with * show False by simp | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3733 | qed | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3734 | |
| 76589 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3735 | lemma irreflp_multp: "transp R \<Longrightarrow> irreflp R \<Longrightarrow> irreflp (multp R)" | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3736 |   by (rule irrefl_mult[of "{(x, y). r x y}" for r,
 | 
| 
1c083e32aed6
stated goals of some lemmas explicitely to prevent silent changes
 desharna parents: 
76570diff
changeset | 3737 | folded transp_trans_eq irreflp_irrefl_eq, simplified, folded multp_def]) | 
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3738 | |
| 74864 | 3739 | instantiation multiset :: (preorder) order begin | 
| 3740 | ||
| 3741 | definition less_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" | |
| 3742 | where "M < N \<longleftrightarrow> multp (<) M N" | |
| 3743 | ||
| 3744 | definition less_eq_multiset :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" | |
| 3745 | where "less_eq_multiset M N \<longleftrightarrow> M < N \<or> M = N" | |
| 3746 | ||
| 3747 | instance | |
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3748 | proof intro_classes | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3749 | fix M N :: "'a multiset" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3750 | show "(M < N) = (M \<le> N \<and> \<not> N \<le> M)" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3751 | unfolding less_eq_multiset_def less_multiset_def | 
| 76749 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 desharna parents: 
76682diff
changeset | 3752 | by (metis irreflp_def irreflp_on_less irreflp_multp transpE transp_on_less transp_multp) | 
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3753 | next | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3754 | fix M :: "'a multiset" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3755 | show "M \<le> M" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3756 | unfolding less_eq_multiset_def | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3757 | by simp | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3758 | next | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3759 | fix M1 M2 M3 :: "'a multiset" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3760 | show "M1 \<le> M2 \<Longrightarrow> M2 \<le> M3 \<Longrightarrow> M1 \<le> M3" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3761 | unfolding less_eq_multiset_def less_multiset_def | 
| 76749 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 desharna parents: 
76682diff
changeset | 3762 | using transp_multp[OF transp_on_less, THEN transpD] | 
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3763 | by blast | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3764 | next | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3765 | fix M N :: "'a multiset" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3766 | show "M \<le> N \<Longrightarrow> N \<le> M \<Longrightarrow> M = N" | 
| 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3767 | unfolding less_eq_multiset_def less_multiset_def | 
| 76749 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 desharna parents: 
76682diff
changeset | 3768 | using transp_multp[OF transp_on_less, THEN transpD] | 
| 
11a24dab1880
strengthened and renamed lemmas preorder.transp_(ge|gr|le|less)
 desharna parents: 
76682diff
changeset | 3769 | using irreflp_multp[OF transp_on_less irreflp_on_less, unfolded irreflp_def, rule_format] | 
| 74865 
b5031a8f7718
added lemmas irreflp_{less,greater} to preorder and {trans,irrefl}_mult{,p} to Multiset
 desharna parents: 
74864diff
changeset | 3770 | by blast | 
| 74864 | 3771 | qed | 
| 3772 | ||
| 3773 | end | |
| 3774 | ||
| 3775 | lemma mset_le_irrefl [elim!]: | |
| 3776 | fixes M :: "'a::preorder multiset" | |
| 3777 | shows "M < M \<Longrightarrow> R" | |
| 3778 | by simp | |
| 3779 | ||
| 80324 | 3780 | lemma wfp_less_multiset[simp]: | 
| 80345 | 3781 |   assumes wf: "wfp ((<) :: ('a :: preorder) \<Rightarrow> 'a \<Rightarrow> bool)"
 | 
| 80324 | 3782 | shows "wfp ((<) :: 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool)" | 
| 79971 
033f90dc441d
redefined wf as an abbreviation for "wf_on UNIV"
 desharna parents: 
79800diff
changeset | 3783 | unfolding less_multiset_def | 
| 80345 | 3784 | using wfp_multp[OF wf] . | 
| 74868 
2741ef11ccf6
added wfP_less to wellorder and wfP_less_multiset
 desharna parents: 
74865diff
changeset | 3785 | |
| 74864 | 3786 | |
| 77049 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3787 | subsubsection \<open>Strict total-order properties\<close> | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3788 | |
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3789 | lemma total_on_mult: | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3790 | assumes "total_on A r" and "trans r" and "\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3791 | shows "total_on B (mult r)" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3792 | proof (rule total_onI) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3793 | fix M1 M2 assume "M1 \<in> B" and "M2 \<in> B" and "M1 \<noteq> M2" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3794 | let ?I = "M1 \<inter># M2" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3795 | show "(M1, M2) \<in> mult r \<or> (M2, M1) \<in> mult r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3796 |   proof (cases "M1 - ?I = {#} \<or> M2 - ?I = {#}")
 | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3797 | case True | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3798 | with \<open>M1 \<noteq> M2\<close> show ?thesis | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3799 | by (metis Diff_eq_empty_iff_mset diff_intersect_left_idem diff_intersect_right_idem | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3800 | subset_implies_mult subset_mset.less_le) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3801 | next | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3802 | case False | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3803 | from assms(1) have "total_on (set_mset (M1 - ?I)) r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3804 | by (meson \<open>M1 \<in> B\<close> assms(3) diff_subset_eq_self set_mset_mono total_on_subset) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3805 | with False obtain greatest1 where | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3806 | greatest1_in: "greatest1 \<in># M1 - ?I" and | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3807 | greatest1_greatest: "\<forall>x \<in># M1 - ?I. greatest1 \<noteq> x \<longrightarrow> (x, greatest1) \<in> r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3808 | using Multiset.bex_greatest_element[to_set, of "M1 - ?I" r] | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3809 | by (metis assms(2) subset_UNIV trans_on_subset) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3810 | |
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3811 | from assms(1) have "total_on (set_mset (M2 - ?I)) r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3812 | by (meson \<open>M2 \<in> B\<close> assms(3) diff_subset_eq_self set_mset_mono total_on_subset) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3813 | with False obtain greatest2 where | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3814 | greatest2_in: "greatest2 \<in># M2 - ?I" and | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3815 | greatest2_greatest: "\<forall>x \<in># M2 - ?I. greatest2 \<noteq> x \<longrightarrow> (x, greatest2) \<in> r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3816 | using Multiset.bex_greatest_element[to_set, of "M2 - ?I" r] | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3817 | by (metis assms(2) subset_UNIV trans_on_subset) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3818 | |
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3819 | have "greatest1 \<noteq> greatest2" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3820 | using greatest1_in \<open>greatest2 \<in># M2 - ?I\<close> | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3821 | by (metis diff_intersect_left_idem diff_intersect_right_idem dual_order.eq_iff in_diff_count | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3822 | in_diff_countE le_add_same_cancel2 less_irrefl zero_le) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3823 | hence "(greatest1, greatest2) \<in> r \<or> (greatest2, greatest1) \<in> r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3824 | using \<open>total_on A r\<close>[unfolded total_on_def, rule_format, of greatest1 greatest2] | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3825 | \<open>M1 \<in> B\<close> \<open>M2 \<in> B\<close> greatest1_in greatest2_in assms(3) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3826 | by (meson in_diffD in_mono) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3827 | thus ?thesis | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3828 | proof (elim disjE) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3829 | assume "(greatest1, greatest2) \<in> r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3830 | have "(?I + (M1 - ?I), ?I + (M2 - ?I)) \<in> mult r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3831 | proof (rule one_step_implies_mult[of "M2 - ?I" "M1 - ?I" r ?I]) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3832 |         show "M2 - ?I \<noteq> {#}"
 | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3833 | using False by force | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3834 | next | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3835 | show "\<forall>k\<in>#M1 - ?I. \<exists>j\<in>#M2 - ?I. (k, j) \<in> r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3836 | using \<open>(greatest1, greatest2) \<in> r\<close> greatest2_in greatest1_greatest | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3837 | by (metis assms(2) transD) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3838 | qed | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3839 | hence "(M1, M2) \<in> mult r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3840 | by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1 | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3841 | subset_mset.inf.cobounded2) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3842 | thus "(M1, M2) \<in> mult r \<or> (M2, M1) \<in> mult r" .. | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3843 | next | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3844 | assume "(greatest2, greatest1) \<in> r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3845 | have "(?I + (M2 - ?I), ?I + (M1 - ?I)) \<in> mult r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3846 | proof (rule one_step_implies_mult[of "M1 - ?I" "M2 - ?I" r ?I]) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3847 |         show "M1 - M1 \<inter># M2 \<noteq> {#}"
 | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3848 | using False by force | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3849 | next | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3850 | show "\<forall>k\<in>#M2 - ?I. \<exists>j\<in>#M1 - ?I. (k, j) \<in> r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3851 | using \<open>(greatest2, greatest1) \<in> r\<close> greatest1_in greatest2_greatest | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3852 | by (metis assms(2) transD) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3853 | qed | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3854 | hence "(M2, M1) \<in> mult r" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3855 | by (metis subset_mset.add_diff_inverse subset_mset.inf.cobounded1 | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3856 | subset_mset.inf.cobounded2) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3857 | thus "(M1, M2) \<in> mult r \<or> (M2, M1) \<in> mult r" .. | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3858 | qed | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3859 | qed | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3860 | qed | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3861 | |
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3862 | lemma total_mult: "total r \<Longrightarrow> trans r \<Longrightarrow> total (mult r)" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3863 | by (rule total_on_mult[of UNIV r UNIV, simplified]) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3864 | |
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3865 | lemma totalp_on_multp: | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3866 | "totalp_on A R \<Longrightarrow> transp R \<Longrightarrow> (\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A) \<Longrightarrow> totalp_on B (multp R)" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3867 |   using total_on_mult[of A "{(x,y). R x y}" B, to_pred]
 | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3868 | by (simp add: multp_def total_on_def totalp_on_def) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3869 | |
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3870 | lemma totalp_multp: "totalp R \<Longrightarrow> transp R \<Longrightarrow> totalp (multp R)" | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3871 | by (rule totalp_on_multp[of UNIV R UNIV, simplified]) | 
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3872 | |
| 
e293216df994
added lemmas total_on_mult, total_mult, totalp_on_multp, and totalp_multp
 desharna parents: 
76755diff
changeset | 3873 | |
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3874 | subsection \<open>Quasi-executable version of the multiset extension\<close> | 
| 63088 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3875 | |
| 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3876 | text \<open> | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3877 | Predicate variants of \<open>mult\<close> and the reflexive closure of \<open>mult\<close>, which are | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3878 | executable whenever the given predicate \<open>P\<close> is. Together with the standard | 
| 67398 | 3879 | code equations for \<open>(\<inter>#\<close>) and \<open>(-\<close>) this should yield quadratic | 
| 74803 | 3880 | (with respect to calls to \<open>P\<close>) implementations of \<open>multp_code\<close> and \<open>multeqp_code\<close>. | 
| 63088 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3881 | \<close> | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3882 | |
| 74803 | 3883 | definition multp_code :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
 | 
| 3884 | "multp_code P N M = | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3885 | (let Z = M \<inter># N; X = M - Z in | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3886 |     X \<noteq> {#} \<and> (let Y = N - Z in (\<forall>y \<in> set_mset Y. \<exists>x \<in> set_mset X. P y x)))"
 | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3887 | |
| 74803 | 3888 | definition multeqp_code :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset \<Rightarrow> bool" where
 | 
| 3889 | "multeqp_code P N M = | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3890 | (let Z = M \<inter># N; X = M - Z; Y = N - Z in | 
| 63088 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3891 | (\<forall>y \<in> set_mset Y. \<exists>x \<in> set_mset X. P y x))" | 
| 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3892 | |
| 74805 | 3893 | lemma multp_code_iff_mult: | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3894 | assumes "irrefl_on (set_mset N \<inter> set_mset M) R" and "trans R" and | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3895 | [simp]: "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R" | 
| 74803 | 3896 | shows "multp_code P N M \<longleftrightarrow> (N, M) \<in> mult R" (is "?L \<longleftrightarrow> ?R") | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3897 | proof - | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3898 | have *: "M \<inter># N + (N - M \<inter># N) = N" "M \<inter># N + (M - M \<inter># N) = M" | 
| 68406 | 3899 |     "(M - M \<inter># N) \<inter># (N - M \<inter># N) = {#}" by (auto simp flip: count_inject)
 | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3900 | show ?thesis | 
| 63088 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3901 | proof | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3902 | assume ?L thus ?R | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3903 | using one_step_implies_mult[of "M - M \<inter># N" "N - M \<inter># N" R "M \<inter># N"] * | 
| 74803 | 3904 | by (auto simp: multp_code_def Let_def) | 
| 63088 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3905 | next | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3906 |     { fix I J K :: "'a multiset" assume "(I + J) \<inter># (I + K) = {#}"
 | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3907 |       then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty)
 | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3908 | } note [dest!] = this | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3909 | assume ?R thus ?L | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3910 | using mult_cancel_max | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3911 | using mult_implies_one_step[OF assms(2), of "N - M \<inter># N" "M - M \<inter># N"] | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3912 | mult_cancel_max[OF assms(2,1)] * by (auto simp: multp_code_def) | 
| 63088 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3913 | qed | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3914 | qed | 
| 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3915 | |
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3916 | lemma multp_code_iff_multp: | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3917 | "irreflp_on (set_mset M \<inter> set_mset N) R \<Longrightarrow> transp R \<Longrightarrow> multp_code R M N \<longleftrightarrow> multp R M N" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3918 | using multp_code_iff_mult[simplified, to_pred, of M N R R] by simp | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3919 | |
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3920 | lemma multp_code_eq_multp: | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3921 | assumes "irreflp R" and "transp R" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3922 | shows "multp_code R = multp R" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3923 | proof (intro ext) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3924 | fix M N | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3925 | show "multp_code R M N = multp R M N" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3926 | proof (rule multp_code_iff_multp) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3927 | from assms show "irreflp_on (set_mset M \<inter> set_mset N) R" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3928 | by (auto intro: irreflp_on_subset) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3929 | next | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3930 | from assms show "transp R" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3931 | by simp | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3932 | qed | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3933 | qed | 
| 74863 
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
 desharna parents: 
74862diff
changeset | 3934 | |
| 74805 | 3935 | lemma multeqp_code_iff_reflcl_mult: | 
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3936 | assumes "irrefl_on (set_mset N \<inter> set_mset M) R" and "trans R" and "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R" | 
| 74803 | 3937 | shows "multeqp_code P N M \<longleftrightarrow> (N, M) \<in> (mult R)\<^sup>=" | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3938 | proof - | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 3939 |   { assume "N \<noteq> M" "M - M \<inter># N = {#}"
 | 
| 68406 | 3940 | then obtain y where "count N y \<noteq> count M y" by (auto simp flip: count_inject) | 
| 64911 | 3941 |     then have "\<exists>y. count M y < count N y" using \<open>M - M \<inter># N = {#}\<close>
 | 
| 68406 | 3942 | by (auto simp flip: count_inject dest!: le_neq_implies_less fun_cong[of _ _ y]) | 
| 63660 
76302202a92d
add monotonicity propertyies of `mult1` and `mult`
 Bertram Felgenhauer <bertram.felgenhauer@uibk.ac.at> parents: 
63560diff
changeset | 3943 | } | 
| 74803 | 3944 | then have "multeqp_code P N M \<longleftrightarrow> multp_code P N M \<or> N = M" | 
| 3945 | by (auto simp: multeqp_code_def multp_code_def Let_def in_diff_count) | |
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3946 | thus ?thesis | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3947 | using multp_code_iff_mult[OF assms] by simp | 
| 63088 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3948 | qed | 
| 
f2177f5d2aed
a quasi-recursive characterization of the multiset order (by Christian Sternagel)
 haftmann parents: 
63060diff
changeset | 3949 | |
| 76611 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3950 | lemma multeqp_code_iff_reflclp_multp: | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3951 | "irreflp_on (set_mset M \<inter> set_mset N) R \<Longrightarrow> transp R \<Longrightarrow> multeqp_code R M N \<longleftrightarrow> (multp R)\<^sup>=\<^sup>= M N" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3952 | using multeqp_code_iff_reflcl_mult[simplified, to_pred, of M N R R] by simp | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3953 | |
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3954 | lemma multeqp_code_eq_reflclp_multp: | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3955 | assumes "irreflp R" and "transp R" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3956 | shows "multeqp_code R = (multp R)\<^sup>=\<^sup>=" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3957 | proof (intro ext) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3958 | fix M N | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3959 | show "multeqp_code R M N \<longleftrightarrow> (multp R)\<^sup>=\<^sup>= M N" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3960 | proof (rule multeqp_code_iff_reflclp_multp) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3961 | from assms show "irreflp_on (set_mset M \<inter> set_mset N) R" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3962 | by (auto intro: irreflp_on_subset) | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3963 | next | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3964 | from assms show "transp R" | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3965 | by simp | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3966 | qed | 
| 
a7d2a7a737b8
Strengthened multiset lemmas w.r.t. irrefl and irreflp
 desharna parents: 
76589diff
changeset | 3967 | qed | 
| 74863 
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
 desharna parents: 
74862diff
changeset | 3968 | |
| 
691131ce4641
added lemmas multp_code_eq_multp and multeqp_code_eq_reflclp_multp
 desharna parents: 
74862diff
changeset | 3969 | |
| 60500 | 3970 | subsubsection \<open>Monotonicity of multiset union\<close> | 
| 10249 | 3971 | |
| 60606 | 3972 | lemma mult1_union: "(B, D) \<in> mult1 r \<Longrightarrow> (C + B, C + D) \<in> mult1 r" | 
| 64076 | 3973 | by (force simp: mult1_def) | 
| 10249 | 3974 | |
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 3975 | lemma union_le_mono2: "B < D \<Longrightarrow> C + B < C + (D::'a::preorder multiset)" | 
| 80095 | 3976 | unfolding less_multiset_def multp_def mult_def | 
| 3977 | by (induction rule: trancl_induct; blast intro: mult1_union trancl_trans) | |
| 10249 | 3978 | |
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 3979 | lemma union_le_mono1: "B < D \<Longrightarrow> B + C < D + (C::'a::preorder multiset)" | 
| 80095 | 3980 | by (metis add.commute union_le_mono2) | 
| 10249 | 3981 | |
| 17161 | 3982 | lemma union_less_mono: | 
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 3983 | fixes A B C D :: "'a::preorder multiset" | 
| 63388 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 3984 | shows "A < C \<Longrightarrow> B < D \<Longrightarrow> A + B < C + D" | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 3985 | by (blast intro!: union_le_mono1 union_le_mono2 less_trans) | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 3986 | |
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 3987 | instantiation multiset :: (preorder) ordered_ab_semigroup_add | 
| 63388 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 3988 | begin | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 3989 | instance | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 3990 | by standard (auto simp add: less_eq_multiset_def intro: union_le_mono2) | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 3991 | end | 
| 15072 | 3992 | |
| 63409 
3f3223b90239
moved lemmas and locales around (with minor incompatibilities)
 blanchet parents: 
63388diff
changeset | 3993 | |
| 60500 | 3994 | 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 | 3995 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 3996 | 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 | 3997 |   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 | 3998 |   and multi_member_last: "x \<in># {# x #}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 3999 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4000 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4001 | definition "ms_strict = mult pair_less" | 
| 37765 | 4002 | 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 | 4003 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4004 | lemma ms_reduction_pair: "reduction_pair (ms_strict, ms_weak)" | 
| 80095 | 4005 | unfolding reduction_pair_def ms_strict_def ms_weak_def pair_less_def | 
| 4006 | by (auto intro: wf_mult1 wf_trancl simp: mult_def) | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4007 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4008 | lemma smsI: | 
| 60495 | 4009 | "(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 | 4010 | unfolding ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4011 | 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 | 4012 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4013 | lemma wmsI: | 
| 60495 | 4014 |   "(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 | 4015 | \<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 | 4016 | unfolding ms_weak_def ms_strict_def | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4017 | 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 | 4018 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4019 | inductive pw_leq | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4020 | where | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4021 |   pw_leq_empty: "pw_leq {#} {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4022 | | 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 | 4023 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4024 | lemma pw_leq_lstep: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4025 |   "(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 | 4026 | 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 | 4027 | |
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4028 | lemma pw_leq_split: | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4029 | assumes "pw_leq X Y" | 
| 60495 | 4030 |   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 | 4031 | using assms | 
| 60606 | 4032 | proof induct | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4033 | 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 | 4034 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4035 | 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 | 4036 | then obtain A B Z where | 
| 58425 | 4037 | [simp]: "X = A + Z" "Y = B + Z" | 
| 60495 | 4038 |       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 | 4039 | by auto | 
| 60606 | 4040 | 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 | 4041 | unfolding pair_leq_def by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4042 | thus ?case | 
| 60606 | 4043 | proof cases | 
| 4044 | case [simp]: 1 | |
| 4045 |     have "{#x#} + X = A + ({#y#}+Z) \<and> {#y#} + Y = B + ({#y#}+Z) \<and>
 | |
| 4046 |       ((set_mset A, set_mset B) \<in> max_strict \<or> (B = {#} \<and> A = {#}))"
 | |
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 4047 | by auto | 
| 60606 | 4048 | thus ?thesis by blast | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4049 | next | 
| 60606 | 4050 | case 2 | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4051 |     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 | 4052 |     have "{#x#} + X = ?A' + Z"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4053 |       "{#y#} + Y = ?B' + Z"
 | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 4054 | by auto | 
| 58425 | 4055 | moreover have | 
| 60495 | 4056 | "(set_mset ?A', set_mset ?B') \<in> max_strict" | 
| 60606 | 4057 | 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 | 4058 | by (auto elim!: max_ext.cases) | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4059 | ultimately show ?thesis by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4060 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4061 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4062 | |
| 58425 | 4063 | lemma | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4064 | assumes pwleq: "pw_leq Z Z'" | 
| 60495 | 4065 | shows ms_strictI: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_strict" | 
| 60606 | 4066 | and ms_weakI1: "(set_mset A, set_mset B) \<in> max_strict \<Longrightarrow> (Z + A, Z' + B) \<in> ms_weak" | 
| 4067 |     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 | 4068 | proof - | 
| 58425 | 4069 | from pw_leq_split[OF pwleq] | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4070 | obtain A' B' Z'' | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4071 | where [simp]: "Z = A' + Z''" "Z' = B' + Z''" | 
| 60495 | 4072 |     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 | 4073 | by blast | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4074 |   {
 | 
| 60495 | 4075 | 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 | 4076 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4077 | 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 | 4078 | proof | 
| 60495 | 4079 | assume max': "(set_mset A', set_mset B') \<in> max_strict" | 
| 4080 | 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 | 4081 | by (auto simp: max_strict_def intro: max_ext_additive) | 
| 58425 | 4082 | thus ?thesis by (rule smsI) | 
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4083 | next | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4084 |       assume [simp]: "A' = {#} \<and> B' = {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4085 | 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 | 4086 | qed | 
| 60606 | 4087 | 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 | 4088 | 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 | 4089 | } | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4090 | from mx_or_empty | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4091 | have "(Z'' + A', Z'' + B') \<in> ms_weak" by (rule wmsI) | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 4092 |   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 | 4093 | qed | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4094 | |
| 39301 | 4095 | 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 | 4096 | and nonempty_plus: "{# x #} + rs \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4097 | and nonempty_single: "{# x #} \<noteq> {#}"
 | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4098 | by auto | 
| 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4099 | |
| 60500 | 4100 | setup \<open> | 
| 60606 | 4101 | let | 
| 74634 | 4102 | fun msetT T = \<^Type>\<open>multiset T\<close>; | 
| 4103 | ||
| 4104 |     fun mk_mset T [] = \<^instantiate>\<open>'a = T in term \<open>{#}\<close>\<close>
 | |
| 4105 |       | mk_mset T [x] = \<^instantiate>\<open>'a = T and x in term \<open>{#x#}\<close>\<close>
 | |
| 4106 | | mk_mset T (x :: xs) = \<^Const>\<open>plus \<open>msetT T\<close> for \<open>mk_mset T [x]\<close> \<open>mk_mset T xs\<close>\<close> | |
| 60606 | 4107 | |
| 60752 | 4108 | fun mset_member_tac ctxt m i = | 
| 60606 | 4109 | if m <= 0 then | 
| 60752 | 4110 |         resolve_tac ctxt @{thms multi_member_this} i ORELSE
 | 
| 4111 |         resolve_tac ctxt @{thms multi_member_last} i
 | |
| 60606 | 4112 | else | 
| 60752 | 4113 |         resolve_tac ctxt @{thms multi_member_skip} i THEN mset_member_tac ctxt (m - 1) i
 | 
| 4114 | ||
| 4115 | fun mset_nonempty_tac ctxt = | |
| 4116 |       resolve_tac ctxt @{thms nonempty_plus} ORELSE'
 | |
| 4117 |       resolve_tac ctxt @{thms nonempty_single}
 | |
| 29125 
d41182a8135c
method "sizechange" proves termination of functions; added more infrastructure for termination proofs
 krauss parents: 
28823diff
changeset | 4118 | |
| 60606 | 4119 | fun regroup_munion_conv ctxt = | 
| 73393 | 4120 | Function_Lib.regroup_conv ctxt \<^const_abbrev>\<open>empty_mset\<close> \<^const_name>\<open>plus\<close> | 
| 60606 | 4121 |         (map (fn t => t RS eq_reflection) (@{thms ac_simps} @ @{thms empty_neutral}))
 | 
| 4122 | ||
| 60752 | 4123 | fun unfold_pwleq_tac ctxt i = | 
| 4124 |       (resolve_tac ctxt @{thms pw_leq_step} i THEN (fn st => unfold_pwleq_tac ctxt (i + 1) st))
 | |
| 4125 |         ORELSE (resolve_tac ctxt @{thms pw_leq_lstep} i)
 | |
| 4126 |         ORELSE (resolve_tac ctxt @{thms pw_leq_empty} i)
 | |
| 60606 | 4127 | |
| 4128 |     val set_mset_simps = [@{thm set_mset_empty}, @{thm set_mset_single}, @{thm set_mset_union},
 | |
| 4129 |                         @{thm Un_insert_left}, @{thm Un_empty_left}]
 | |
| 4130 | in | |
| 4131 | ScnpReconstruct.multiset_setup (ScnpReconstruct.Multiset | |
| 4132 |     {
 | |
| 4133 | msetT=msetT, mk_mset=mk_mset, mset_regroup_conv=regroup_munion_conv, | |
| 4134 | mset_member_tac=mset_member_tac, mset_nonempty_tac=mset_nonempty_tac, | |
| 4135 | mset_pwleq_tac=unfold_pwleq_tac, set_of_simps=set_mset_simps, | |
| 4136 |       smsI'= @{thm ms_strictI}, wmsI2''= @{thm ms_weakI2}, wmsI1= @{thm ms_weakI1},
 | |
| 60752 | 4137 |       reduction_pair = @{thm ms_reduction_pair}
 | 
| 60606 | 4138 | }) | 
| 4139 | end | |
| 60500 | 4140 | \<close> | 
| 4141 | ||
| 4142 | ||
| 4143 | 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 | 4144 | |
| 39302 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 nipkow parents: 
39301diff
changeset | 4145 | 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 | 4146 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4147 | 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 | 4148 | 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 | 4149 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4150 | 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 | 4151 | 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 | 4152 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4153 | 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 | 4154 | 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 | 4155 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4156 | lemmas union_ac = union_assoc union_commute union_lcomm add_mset_commute | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4157 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4158 | 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 | 4159 | 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 | 4160 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4161 | 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 | 4162 | 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 | 4163 | |
| 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4164 | lemma multi_union_self_other_eq: "(A::'a multiset) + X = A + Y \<Longrightarrow> X = Y" | 
| 59557 | 4165 | 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 | 4166 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4167 | lemma mset_subset_trans: "(M::'a multiset) \<subset># K \<Longrightarrow> K \<subset># N \<Longrightarrow> M \<subset># N" | 
| 60397 
f8a513fedb31
Renaming multiset operators < ~> <#,...
 Mathias Fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
59999diff
changeset | 4168 | by (fact subset_mset.less_trans) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4169 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 4170 | 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 | 4171 | by (fact subset_mset.inf.commute) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4172 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 4173 | 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 | 4174 | by (fact subset_mset.inf.assoc [symmetric]) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4175 | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 4176 | 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 | 4177 | by (fact subset_mset.inf.left_commute) | 
| 35268 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4178 | |
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4179 | lemmas multiset_inter_ac = | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4180 | multiset_inter_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4181 | multiset_inter_assoc | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4182 | multiset_inter_left_commute | 
| 
04673275441a
switched notations for pointwise and multiset order
 haftmann parents: 
35028diff
changeset | 4183 | |
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 4184 | lemma mset_le_not_refl: "\<not> M < (M::'a::preorder multiset)" | 
| 63388 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 4185 | by (fact less_irrefl) | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 4186 | |
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 4187 | lemma mset_le_trans: "K < M \<Longrightarrow> M < N \<Longrightarrow> K < (N::'a::preorder multiset)" | 
| 63388 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 4188 | by (fact less_trans) | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 4189 | |
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 4190 | lemma mset_le_not_sym: "M < N \<Longrightarrow> \<not> N < (M::'a::preorder multiset)" | 
| 63388 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 4191 | by (fact less_not_sym) | 
| 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 4192 | |
| 63410 
9789ccc2a477
more instantiations for multiset
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63409diff
changeset | 4193 | lemma mset_le_asym: "M < N \<Longrightarrow> (\<not> P \<Longrightarrow> N < (M::'a::preorder multiset)) \<Longrightarrow> P" | 
| 63388 
a095acd4cfbf
instantiate multiset with multiset ordering
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63360diff
changeset | 4194 | by (fact less_asym) | 
| 34943 
e97b22500a5c
cleanup of Multiset.thy: less duplication, tuned and simplified a couple of proofs, less historical organization of sections, conversion from associations lists to multisets, rudimentary code generation
 haftmann parents: 
33102diff
changeset | 4195 | |
| 60500 | 4196 | declaration \<open> | 
| 60606 | 4197 | let | 
| 4198 | fun multiset_postproc _ maybe_name all_values (T as Type (_, [elem_T])) (Const _ $ t') = | |
| 4199 | let | |
| 4200 | val (maybe_opt, ps) = | |
| 4201 | Nitpick_Model.dest_plain_fun t' | |
| 67398 | 4202 | ||> (~~) | 
| 60606 | 4203 | ||> map (apsnd (snd o HOLogic.dest_number)) | 
| 4204 | fun elems_for t = | |
| 67398 | 4205 | (case AList.lookup (=) ps t of | 
| 60606 | 4206 | SOME n => replicate n t | 
| 4207 | | NONE => [Const (maybe_name, elem_T --> elem_T) $ t]) | |
| 4208 | in | |
| 4209 | (case maps elems_for (all_values elem_T) @ | |
| 61333 | 4210 | (if maybe_opt then [Const (Nitpick_Model.unrep_mixfix (), elem_T)] else []) of | 
| 74634 | 4211 | [] => \<^Const>\<open>Groups.zero T\<close> | 
| 4212 | | ts => foldl1 (fn (s, t) => \<^Const>\<open>add_mset elem_T for s t\<close>) ts) | |
| 60606 | 4213 | end | 
| 4214 | | multiset_postproc _ _ _ _ t = t | |
| 69593 | 4215 | in Nitpick_Model.register_term_postprocessor \<^typ>\<open>'a multiset\<close> multiset_postproc end | 
| 60500 | 4216 | \<close> | 
| 4217 | ||
| 4218 | ||
| 4219 | 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 | 4220 | |
| 60515 | 4221 | code_datatype mset | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4222 | |
| 60606 | 4223 | lemma [code]: "{#} = mset []"
 | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4224 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4225 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4226 | lemma [code]: "add_mset x (mset xs) = mset (x # xs)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4227 | by simp | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4228 | |
| 63195 | 4229 | lemma [code]: "Multiset.is_empty (mset xs) \<longleftrightarrow> List.null xs" | 
| 4230 | by (simp add: Multiset.is_empty_def List.null_def) | |
| 4231 | ||
| 60606 | 4232 | 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 | 4233 | by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4234 | |
| 60606 | 4235 | lemma [code]: "image_mset f (mset xs) = mset (map f xs)" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4236 | by simp | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4237 | |
| 60606 | 4238 | lemma [code]: "filter_mset f (mset xs) = mset (filter f xs)" | 
| 69442 | 4239 | by simp | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4240 | |
| 60606 | 4241 | lemma [code]: "mset xs - mset ys = mset (fold remove1 ys xs)" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4242 | by (rule sym, induct ys arbitrary: xs) (simp_all add: diff_add diff_right_commute diff_diff_add) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4243 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4244 | lemma [code]: | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 4245 | "mset xs \<inter># mset ys = | 
| 60515 | 4246 | 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 | 4247 | 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 | 4248 | proof - | 
| 60515 | 4249 | 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 | 4250 | if x \<in> set ys then (remove1 x ys, x # zs) else (ys, zs)) xs (ys, zs))) = | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 4251 | (mset xs \<inter># mset ys) + mset zs" | 
| 51623 | 4252 | by (induct xs arbitrary: ys) | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 4253 | (auto simp add: inter_add_right1 inter_add_right2 ac_simps) | 
| 51623 | 4254 | then show ?thesis by simp | 
| 4255 | qed | |
| 4256 | ||
| 4257 | lemma [code]: | |
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 4258 | "mset xs \<union># mset ys = | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61378diff
changeset | 4259 | mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, [])))" | 
| 51623 | 4260 | proof - | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61378diff
changeset | 4261 | have "\<And>zs. mset (case_prod append (fold (\<lambda>x (ys, zs). (remove1 x ys, x # zs)) xs (ys, zs))) = | 
| 63919 
9aed2da07200
# after multiset intersection and union symbol
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63908diff
changeset | 4262 | (mset xs \<union># mset ys) + mset zs" | 
| 51623 | 4263 | 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 | 4264 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4265 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4266 | |
| 59813 | 4267 | 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 | 4268 | |
| 60606 | 4269 | 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 | 4270 | proof - | 
| 60515 | 4271 | 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 | 4272 | by (induct xs) simp_all | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4273 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4274 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4275 | |
| 60515 | 4276 | declare set_mset_mset [code] | 
| 4277 | ||
| 4278 | 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 | 4279 | |
| 61585 | 4280 | lemma [code]: \<comment> \<open>not very efficient, but representation-ignorant!\<close> | 
| 60515 | 4281 | "mset_set A = mset (sorted_list_of_set A)" | 
| 80095 | 4282 | by (metis mset_sorted_list_of_multiset sorted_list_of_mset_set) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4283 | |
| 60515 | 4284 | declare size_mset [code] | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4285 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4286 | fun subset_eq_mset_impl :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool option" where | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4287 | "subset_eq_mset_impl [] ys = Some (ys \<noteq> [])" | 
| 67398 | 4288 | | "subset_eq_mset_impl (Cons x xs) ys = (case List.extract ((=) x) ys of | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4289 | None \<Rightarrow> None | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4290 | | Some (ys1,_,ys2) \<Rightarrow> subset_eq_mset_impl xs (ys1 @ ys2))" | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4291 | |
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4292 | lemma subset_eq_mset_impl: "(subset_eq_mset_impl xs ys = None \<longleftrightarrow> \<not> mset xs \<subseteq># mset ys) \<and> | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4293 | (subset_eq_mset_impl xs ys = Some True \<longleftrightarrow> mset xs \<subset># mset ys) \<and> | 
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4294 | (subset_eq_mset_impl xs ys = Some False \<longrightarrow> mset xs = mset ys)" | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4295 | proof (induct xs arbitrary: ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4296 | case (Nil ys) | 
| 64076 | 4297 | show ?case by (auto simp: subset_mset.zero_less_iff_neq_zero) | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4298 | next | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4299 | case (Cons x xs ys) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4300 | show ?case | 
| 67398 | 4301 | proof (cases "List.extract ((=) x) ys") | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4302 | case None | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4303 | 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 | 4304 |     {
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 4305 | assume "mset (x # xs) \<subseteq># mset ys" | 
| 60495 | 4306 | 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 | 4307 | } note nle = this | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4308 | moreover | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4309 |     {
 | 
| 62430 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 4310 | assume "mset (x # xs) \<subset># mset ys" | 
| 
9527ff088c15
more succint formulation of membership for multisets, similar to lists;
 haftmann parents: 
62390diff
changeset | 4311 | 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 | 4312 | 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 | 4313 | } | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4314 | 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 | 4315 | next | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4316 | case (Some res) | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4317 | 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 | 4318 | note Some = Some[unfolded res] | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4319 | from extract_SomeE[OF Some] have "ys = ys1 @ x # ys2" by simp | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4320 | hence id: "mset ys = add_mset x (mset (ys1 @ ys2))" | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 4321 | by auto | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4322 | show ?thesis unfolding subset_eq_mset_impl.simps | 
| 80095 | 4323 | by (simp add: Some id Cons) | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4324 | qed | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4325 | qed | 
| 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4326 | |
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4327 | lemma [code]: "mset xs \<subseteq># mset ys \<longleftrightarrow> subset_eq_mset_impl xs ys \<noteq> None" | 
| 80095 | 4328 | by (simp add: subset_eq_mset_impl) | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4329 | |
| 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4330 | lemma [code]: "mset xs \<subset># mset ys \<longleftrightarrow> subset_eq_mset_impl xs ys = Some True" | 
| 80095 | 4331 | using subset_eq_mset_impl by blast | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4332 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4333 | instantiation multiset :: (equal) equal | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4334 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4335 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4336 | definition | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4337 | [code del]: "HOL.equal A (B :: 'a multiset) \<longleftrightarrow> A = B" | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4338 | lemma [code]: "HOL.equal (mset xs) (mset ys) \<longleftrightarrow> subset_eq_mset_impl xs ys = Some False" | 
| 55808 
488c3e8282c8
added Rene Thiemann's patch for the nonterminating equality/subset test code for multisets
 nipkow parents: 
55565diff
changeset | 4339 | unfolding equal_multiset_def | 
| 63310 
caaacf37943f
normalising multiset theorem names
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63290diff
changeset | 4340 | using subset_eq_mset_impl[of xs ys] by (cases "subset_eq_mset_impl xs ys", auto) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4341 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4342 | instance | 
| 60678 | 4343 | by standard (simp add: equal_multiset_def) | 
| 4344 | ||
| 37169 
f69efa106feb
make Nitpick "show_all" option behave less surprisingly
 blanchet parents: 
37107diff
changeset | 4345 | end | 
| 49388 | 4346 | |
| 66313 | 4347 | declare sum_mset_sum_list [code] | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4348 | |
| 63830 | 4349 | lemma [code]: "prod_mset (mset xs) = fold times xs 1" | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4350 | proof - | 
| 63830 | 4351 | have "\<And>x. fold times xs x = prod_mset (mset xs) * x" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4352 | by (induct xs) (simp_all add: ac_simps) | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4353 | then show ?thesis by simp | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4354 | qed | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4355 | |
| 60500 | 4356 | text \<open> | 
| 69593 | 4357 | Exercise for the casual reader: add implementations for \<^term>\<open>(\<le>)\<close> | 
| 4358 | and \<^term>\<open>(<)\<close> (multiset order). | |
| 60500 | 4359 | \<close> | 
| 4360 | ||
| 4361 | text \<open>Quickcheck generators\<close> | |
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4362 | |
| 72607 | 4363 | context | 
| 4364 | includes term_syntax | |
| 4365 | begin | |
| 4366 | ||
| 4367 | definition | |
| 61076 | 4368 | 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 | 4369 | \<Rightarrow> 'a multiset \<times> (unit \<Rightarrow> Code_Evaluation.term)" where | 
| 60515 | 4370 |   [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 | 4371 | |
| 72607 | 4372 | end | 
| 4373 | ||
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4374 | instantiation multiset :: (random) random | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4375 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4376 | |
| 72581 | 4377 | context | 
| 4378 | includes state_combinator_syntax | |
| 4379 | begin | |
| 4380 | ||
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4381 | definition | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4382 | "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 | 4383 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4384 | instance .. | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4385 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4386 | end | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4387 | |
| 72581 | 4388 | end | 
| 51600 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4389 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4390 | instantiation multiset :: (full_exhaustive) full_exhaustive | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4391 | begin | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4392 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4393 | 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 | 4394 | where | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4395 | "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 | 4396 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4397 | instance .. | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4398 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4399 | end | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4400 | |
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4401 | hide_const (open) msetify | 
| 
197e25f13f0c
default implementation of multisets by list with reasonable coverage of operations on multisets
 haftmann parents: 
51599diff
changeset | 4402 | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4403 | |
| 60500 | 4404 | subsection \<open>BNF setup\<close> | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4405 | |
| 57966 | 4406 | definition rel_mset where | 
| 60515 | 4407 | "rel_mset R X Y \<longleftrightarrow> (\<exists>xs ys. mset xs = X \<and> mset ys = Y \<and> list_all2 R xs ys)" | 
| 4408 | ||
| 4409 | lemma mset_zip_take_Cons_drop_twice: | |
| 57966 | 4410 | assumes "length xs = length ys" "j \<le> length xs" | 
| 60515 | 4411 | shows "mset (zip (take j xs @ x # drop j xs) (take j ys @ y # drop j ys)) = | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4412 | add_mset (x,y) (mset (zip xs ys))" | 
| 60606 | 4413 | using assms | 
| 57966 | 4414 | proof (induct xs ys arbitrary: x y j rule: list_induct2) | 
| 4415 | case Nil | |
| 4416 | thus ?case | |
| 4417 | by simp | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4418 | next | 
| 57966 | 4419 | case (Cons x xs y ys) | 
| 4420 | thus ?case | |
| 4421 | proof (cases "j = 0") | |
| 4422 | case True | |
| 4423 | thus ?thesis | |
| 4424 | by simp | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4425 | next | 
| 57966 | 4426 | case False | 
| 4427 | then obtain k where k: "j = Suc k" | |
| 60678 | 4428 | by (cases j) simp | 
| 57966 | 4429 | hence "k \<le> length xs" | 
| 4430 | using Cons.prems by auto | |
| 60515 | 4431 | hence "mset (zip (take k xs @ x # drop k xs) (take k ys @ y # drop k ys)) = | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4432 | add_mset (x,y) (mset (zip xs ys))" | 
| 57966 | 4433 | by (rule Cons.hyps(2)) | 
| 4434 | thus ?thesis | |
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 4435 | unfolding k by auto | 
| 58425 | 4436 | qed | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4437 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4438 | |
| 60515 | 4439 | lemma ex_mset_zip_left: | 
| 4440 | assumes "length xs = length ys" "mset xs' = mset xs" | |
| 4441 | shows "\<exists>ys'. length ys' = length xs' \<and> mset (zip xs' ys') = mset (zip xs ys)" | |
| 58425 | 4442 | using assms | 
| 57966 | 4443 | proof (induct xs ys arbitrary: xs' rule: list_induct2) | 
| 4444 | case Nil | |
| 4445 | thus ?case | |
| 4446 | by auto | |
| 4447 | next | |
| 4448 | case (Cons x xs y ys xs') | |
| 4449 | obtain j where j_len: "j < length xs'" and nth_j: "xs' ! j = x" | |
| 60515 | 4450 | by (metis Cons.prems in_set_conv_nth list.set_intros(1) mset_eq_setD) | 
| 58425 | 4451 | |
| 63040 | 4452 | define xsa where "xsa = take j xs' @ drop (Suc j) xs'" | 
| 60515 | 4453 |   have "mset xs' = {#x#} + mset xsa"
 | 
| 57966 | 4454 | unfolding xsa_def using j_len nth_j | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4455 | by (metis Cons_nth_drop_Suc union_mset_add_mset_right add_mset_remove_trivial add_diff_cancel_left' | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4456 | append_take_drop_id mset.simps(2) mset_append) | 
| 60515 | 4457 | hence ms_x: "mset xsa = mset xs" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4458 | by (simp add: Cons.prems) | 
| 57966 | 4459 | then obtain ysa where | 
| 60515 | 4460 | len_a: "length ysa = length xsa" and ms_a: "mset (zip xsa ysa) = mset (zip xs ys)" | 
| 57966 | 4461 | using Cons.hyps(2) by blast | 
| 4462 | ||
| 63040 | 4463 | define ys' where "ys' = take j ysa @ y # drop j ysa" | 
| 57966 | 4464 | have xs': "xs' = take j xsa @ x # drop j xsa" | 
| 4465 | 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 | 4466 | by (metis append_eq_append_conv append_take_drop_id diff_Suc_Suc Cons_nth_drop_Suc length_Cons | 
| 60515 | 4467 | length_drop size_mset) | 
| 57966 | 4468 | have j_len': "j \<le> length xsa" | 
| 4469 | using j_len xs' xsa_def | |
| 4470 | by (metis add_Suc_right append_take_drop_id length_Cons length_append less_eq_Suc_le not_less) | |
| 4471 | have "length ys' = length xs'" | |
| 4472 | unfolding ys'_def using Cons.prems len_a ms_x | |
| 60515 | 4473 | by (metis add_Suc_right append_take_drop_id length_Cons length_append mset_eq_length) | 
| 4474 | moreover have "mset (zip xs' ys') = mset (zip (x # xs) (y # ys))" | |
| 57966 | 4475 | unfolding xs' ys'_def | 
| 60515 | 4476 | by (rule trans[OF mset_zip_take_Cons_drop_twice]) | 
| 63794 
bcec0534aeea
clean argument of simp add
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63793diff
changeset | 4477 | (auto simp: len_a ms_a j_len') | 
| 57966 | 4478 | ultimately show ?case | 
| 4479 | by blast | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4480 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4481 | |
| 57966 | 4482 | lemma list_all2_reorder_left_invariance: | 
| 60515 | 4483 | assumes rel: "list_all2 R xs ys" and ms_x: "mset xs' = mset xs" | 
| 4484 | shows "\<exists>ys'. list_all2 R xs' ys' \<and> mset ys' = mset ys" | |
| 57966 | 4485 | proof - | 
| 4486 | have len: "length xs = length ys" | |
| 4487 | using rel list_all2_conv_all_nth by auto | |
| 4488 | obtain ys' where | |
| 60515 | 4489 | len': "length xs' = length ys'" and ms_xy: "mset (zip xs' ys') = mset (zip xs ys)" | 
| 4490 | using len ms_x by (metis ex_mset_zip_left) | |
| 57966 | 4491 | have "list_all2 R xs' ys'" | 
| 60515 | 4492 | using assms(1) len' ms_xy unfolding list_all2_iff by (blast dest: mset_eq_setD) | 
| 4493 | moreover have "mset ys' = mset ys" | |
| 4494 | using len len' ms_xy map_snd_zip mset_map by metis | |
| 57966 | 4495 | ultimately show ?thesis | 
| 4496 | by blast | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4497 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4498 | |
| 60515 | 4499 | lemma ex_mset: "\<exists>xs. mset xs = X" | 
| 4500 | 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 | 4501 | |
| 73301 | 4502 | inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool" 
 | 
| 62324 | 4503 | where | 
| 4504 |   "pred_mset P {#}"
 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4505 | | "\<lbrakk>P a; pred_mset P M\<rbrakk> \<Longrightarrow> pred_mset P (add_mset a M)" | 
| 62324 | 4506 | |
| 73301 | 4507 | lemma pred_mset_iff: \<comment> \<open>TODO: alias for \<^const>\<open>Multiset.Ball\<close>\<close> | 
| 4508 | \<open>pred_mset P M \<longleftrightarrow> Multiset.Ball M P\<close> (is \<open>?P \<longleftrightarrow> ?Q\<close>) | |
| 4509 | proof | |
| 4510 | assume ?P | |
| 4511 | then show ?Q by induction simp_all | |
| 4512 | next | |
| 4513 | assume ?Q | |
| 4514 | then show ?P | |
| 4515 | by (induction M) (auto intro: pred_mset.intros) | |
| 4516 | qed | |
| 4517 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4518 | bnf "'a multiset" | 
| 57966 | 4519 | map: image_mset | 
| 60495 | 4520 | sets: set_mset | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4521 | bd: natLeq | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4522 |   wits: "{#}"
 | 
| 57966 | 4523 | rel: rel_mset | 
| 62324 | 4524 | pred: pred_mset | 
| 57966 | 4525 | proof - | 
| 60606 | 4526 | show "image_mset (g \<circ> f) = image_mset g \<circ> image_mset f" for f g | 
| 59813 | 4527 | unfolding comp_def by (rule ext) (simp add: comp_def image_mset.compositionality) | 
| 60606 | 4528 | 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 | 4529 | by (induct X) simp_all | 
| 57966 | 4530 | show "card_order natLeq" | 
| 4531 | by (rule natLeq_card_order) | |
| 4532 | show "BNF_Cardinal_Arithmetic.cinfinite natLeq" | |
| 4533 | by (rule natLeq_cinfinite) | |
| 75624 | 4534 | show "regularCard natLeq" | 
| 4535 | by (rule regularCard_natLeq) | |
| 4536 | show "ordLess2 (card_of (set_mset X)) natLeq" for X | |
| 57966 | 4537 | by transfer | 
| 75624 | 4538 | (auto simp: finite_iff_ordLess_natLeq[symmetric]) | 
| 60606 | 4539 | show "rel_mset R OO rel_mset S \<le> rel_mset (R OO S)" for R S | 
| 57966 | 4540 | unfolding rel_mset_def[abs_def] OO_def | 
| 80095 | 4541 | by (smt (verit, ccfv_SIG) list_all2_reorder_left_invariance list_all2_trans predicate2I) | 
| 60606 | 4542 | show "rel_mset R = | 
| 62324 | 4543 |     (\<lambda>x y. \<exists>z. set_mset z \<subseteq> {(x, y). R x y} \<and>
 | 
| 4544 | image_mset fst z = x \<and> image_mset snd z = y)" for R | |
| 4545 | unfolding rel_mset_def[abs_def] | |
| 80095 | 4546 | by (metis (no_types, lifting) ex_mset list.in_rel mem_Collect_eq mset_map set_mset_mset) | 
| 62324 | 4547 | show "pred_mset P = (\<lambda>x. Ball (set_mset x) P)" for P | 
| 73301 | 4548 | by (simp add: fun_eq_iff pred_mset_iff) | 
| 80095 | 4549 | qed auto | 
| 57966 | 4550 | |
| 73301 | 4551 | inductive rel_mset' :: \<open>('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> 'b multiset \<Rightarrow> bool\<close>
 | 
| 60606 | 4552 | where | 
| 57966 | 4553 |   Zero[intro]: "rel_mset' R {#} {#}"
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4554 | | Plus[intro]: "\<lbrakk>R a b; rel_mset' R M N\<rbrakk> \<Longrightarrow> rel_mset' R (add_mset a M) (add_mset b N)" | 
| 57966 | 4555 | |
| 4556 | lemma rel_mset_Zero: "rel_mset R {#} {#}"
 | |
| 4557 | 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 | 4558 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4559 | declare multiset.count[simp] | 
| 73270 | 4560 | declare count_Abs_multiset[simp] | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4561 | declare multiset.count_inverse[simp] | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4562 | |
| 57966 | 4563 | lemma rel_mset_Plus: | 
| 60606 | 4564 | assumes ab: "R a b" | 
| 4565 | and MN: "rel_mset R M N" | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4566 | shows "rel_mset R (add_mset a M) (add_mset b N)" | 
| 60606 | 4567 | proof - | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4568 | have "\<exists>ya. add_mset a (image_mset fst y) = image_mset fst ya \<and> | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4569 | add_mset b (image_mset snd y) = image_mset snd ya \<and> | 
| 60606 | 4570 |     set_mset ya \<subseteq> {(x, y). R x y}"
 | 
| 4571 |     if "R a b" and "set_mset y \<subseteq> {(x, y). R x y}" for y
 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4572 | using that by (intro exI[of _ "add_mset (a,b) y"]) auto | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4573 | thus ?thesis | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4574 | using assms | 
| 57966 | 4575 | 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 | 4576 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4577 | |
| 60606 | 4578 | lemma rel_mset'_imp_rel_mset: "rel_mset' R M N \<Longrightarrow> rel_mset R M N" | 
| 60678 | 4579 | by (induct rule: rel_mset'.induct) (auto simp: rel_mset_Zero rel_mset_Plus) | 
| 57966 | 4580 | |
| 60606 | 4581 | lemma rel_mset_size: "rel_mset R M N \<Longrightarrow> size M = size N" | 
| 60678 | 4582 | 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 | 4583 | |
| 73594 | 4584 | lemma rel_mset_Zero_iff [simp]: | 
| 4585 |   shows "rel_mset rel {#} Y \<longleftrightarrow> Y = {#}" and "rel_mset rel X {#} \<longleftrightarrow> X = {#}"
 | |
| 4586 | by (auto simp add: rel_mset_Zero dest: rel_mset_size) | |
| 4587 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4588 | lemma multiset_induct2[case_names empty addL addR]: | 
| 60678 | 4589 |   assumes empty: "P {#} {#}"
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4590 | and addL: "\<And>a M N. P M N \<Longrightarrow> P (add_mset a M) N" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4591 | and addR: "\<And>a M N. P M N \<Longrightarrow> P M (add_mset a N)" | 
| 60678 | 4592 | shows "P M N" | 
| 80095 | 4593 | by (induct N rule: multiset_induct; induct M rule: multiset_induct) (auto simp: assms) | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4594 | |
| 59949 | 4595 | lemma multiset_induct2_size[consumes 1, case_names empty add]: | 
| 60606 | 4596 | assumes c: "size M = size N" | 
| 4597 |     and empty: "P {#} {#}"
 | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4598 | and add: "\<And>a b M N a b. P M N \<Longrightarrow> P (add_mset a M) (add_mset b N)" | 
| 60606 | 4599 | shows "P M N" | 
| 60678 | 4600 | using c | 
| 4601 | proof (induct M arbitrary: N rule: measure_induct_rule[of size]) | |
| 60606 | 4602 | case (less M) | 
| 4603 | show ?case | |
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4604 |   proof(cases "M = {#}")
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4605 |     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 | 4606 | thus ?thesis using True empty by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4607 | next | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4608 | case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4609 |     have "N \<noteq> {#}" using False less.prems by auto
 | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4610 | then obtain N1 b where N: "N = add_mset b N1" by (metis multi_nonempty_split) | 
| 59949 | 4611 | 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 | 4612 | 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 | 4613 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4614 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4615 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4616 | lemma msed_map_invL: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4617 | assumes "image_mset f (add_mset a M) = N" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4618 | shows "\<exists>N1. N = add_mset (f a) N1 \<and> image_mset f M = N1" | 
| 60606 | 4619 | proof - | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4620 | have "f a \<in># N" | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4621 | using assms multiset.set_map[of f "add_mset a M"] by auto | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4622 | then obtain N1 where N: "N = add_mset (f a) N1" using multi_member_split by metis | 
| 57966 | 4623 | 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 | 4624 | thus ?thesis using N by blast | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4625 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4626 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4627 | lemma msed_map_invR: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4628 | assumes "image_mset f M = add_mset b N" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4629 | shows "\<exists>M1 a. M = add_mset a M1 \<and> f a = b \<and> image_mset f M1 = N" | 
| 60606 | 4630 | proof - | 
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4631 | obtain a where a: "a \<in># M" and fa: "f a = b" | 
| 60606 | 4632 | 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 | 4633 | by (metis image_iff union_single_eq_member) | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4634 | then obtain M1 where M: "M = add_mset a M1" using multi_member_split by metis | 
| 57966 | 4635 | 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 | 4636 | thus ?thesis using M fa by blast | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4637 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4638 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4639 | lemma msed_rel_invL: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4640 | assumes "rel_mset R (add_mset a M) N" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4641 | shows "\<exists>N1 b. N = add_mset b N1 \<and> R a b \<and> rel_mset R M N1" | 
| 60606 | 4642 | proof - | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4643 | obtain K where KM: "image_mset fst K = add_mset a M" | 
| 60606 | 4644 |     and KN: "image_mset snd K = N" and sK: "set_mset K \<subseteq> {(a, b). R a b}"
 | 
| 4645 | using assms | |
| 4646 | unfolding multiset.rel_compp_Grp Grp_def by auto | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4647 | obtain K1 ab where K: "K = add_mset ab K1" and a: "fst ab = a" | 
| 60606 | 4648 | and K1M: "image_mset fst K1 = M" using msed_map_invR[OF KM] by auto | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4649 | obtain N1 where N: "N = add_mset (snd ab) N1" and K1N1: "image_mset snd K1 = N1" | 
| 60606 | 4650 | 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 | 4651 | have Rab: "R a (snd ab)" using sK a unfolding K by auto | 
| 57966 | 4652 | have "rel_mset R M N1" using sK K1M K1N1 | 
| 60606 | 4653 | 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 | 4654 | thus ?thesis using N Rab by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4655 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4656 | |
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4657 | lemma msed_rel_invR: | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4658 | assumes "rel_mset R M (add_mset b N)" | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4659 | shows "\<exists>M1 a. M = add_mset a M1 \<and> R a b \<and> rel_mset R M1 N" | 
| 60606 | 4660 | proof - | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4661 | obtain K where KN: "image_mset snd K = add_mset b N" | 
| 60606 | 4662 |     and KM: "image_mset fst K = M" and sK: "set_mset K \<subseteq> {(a, b). R a b}"
 | 
| 4663 | using assms | |
| 4664 | unfolding multiset.rel_compp_Grp Grp_def by auto | |
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4665 | obtain K1 ab where K: "K = add_mset ab K1" and b: "snd ab = b" | 
| 60606 | 4666 | and K1N: "image_mset snd K1 = N" using msed_map_invR[OF KN] by auto | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4667 | obtain M1 where M: "M = add_mset (fst ab) M1" and K1M1: "image_mset fst K1 = M1" | 
| 60606 | 4668 | 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 | 4669 | have Rab: "R (fst ab) b" using sK b unfolding K by auto | 
| 57966 | 4670 | have "rel_mset R M1 N" using sK K1N K1M1 | 
| 60606 | 4671 | 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 | 4672 | thus ?thesis using M Rab by auto | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4673 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4674 | |
| 57966 | 4675 | lemma rel_mset_imp_rel_mset': | 
| 60606 | 4676 | assumes "rel_mset R M N" | 
| 4677 | shows "rel_mset' R M N" | |
| 59949 | 4678 | 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 | 4679 | case (less M) | 
| 59949 | 4680 | 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 | 4681 | show ?case | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4682 |   proof(cases "M = {#}")
 | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4683 |     case True hence "N = {#}" using c by simp
 | 
| 57966 | 4684 | 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 | 4685 | next | 
| 63793 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4686 | case False then obtain M1 a where M: "M = add_mset a M1" by (metis multi_nonempty_split) | 
| 
e68a0b651eb5
add_mset constructor in multisets
 fleury <Mathias.Fleury@mpi-inf.mpg.de> parents: 
63689diff
changeset | 4687 | obtain N1 b where N: "N = add_mset b N1" and R: "R a b" and ms: "rel_mset R M1 N1" | 
| 60606 | 4688 | using msed_rel_invL[OF less.prems[unfolded M]] by auto | 
| 57966 | 4689 | have "rel_mset' R M1 N1" using less.hyps[of M1 N1] ms unfolding M by simp | 
| 4690 | 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 | 4691 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4692 | qed | 
| 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4693 | |
| 60606 | 4694 | lemma rel_mset_rel_mset': "rel_mset R M N = rel_mset' R M N" | 
| 60678 | 4695 | using rel_mset_imp_rel_mset' rel_mset'_imp_rel_mset by auto | 
| 57966 | 4696 | |
| 69593 | 4697 | text \<open>The main end product for \<^const>\<open>rel_mset\<close>: inductive characterization:\<close> | 
| 61337 | 4698 | lemmas rel_mset_induct[case_names empty add, induct pred: rel_mset] = | 
| 60606 | 4699 | 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 | 4700 | |
| 56656 | 4701 | |
| 60500 | 4702 | subsection \<open>Size setup\<close> | 
| 56656 | 4703 | |
| 67332 | 4704 | lemma size_multiset_o_map: "size_multiset g \<circ> image_mset f = size_multiset (g \<circ> f)" | 
| 65547 | 4705 | apply (rule ext) | 
| 4706 | subgoal for x by (induct x) auto | |
| 4707 | done | |
| 56656 | 4708 | |
| 60500 | 4709 | setup \<open> | 
| 69593 | 4710 | BNF_LFP_Size.register_size_global \<^type_name>\<open>multiset\<close> \<^const_name>\<open>size_multiset\<close> | 
| 62082 | 4711 |     @{thm size_multiset_overloaded_def}
 | 
| 60606 | 4712 |     @{thms size_multiset_empty size_multiset_single size_multiset_union size_empty size_single
 | 
| 4713 | size_union} | |
| 67332 | 4714 |     @{thms size_multiset_o_map}
 | 
| 60500 | 4715 | \<close> | 
| 56656 | 4716 | |
| 65547 | 4717 | subsection \<open>Lemmas about Size\<close> | 
| 4718 | ||
| 4719 | lemma size_mset_SucE: "size A = Suc n \<Longrightarrow> (\<And>a B. A = {#a#} + B \<Longrightarrow> size B = n \<Longrightarrow> P) \<Longrightarrow> P"
 | |
| 4720 | by (cases A) (auto simp add: ac_simps) | |
| 4721 | ||
| 4722 | lemma size_Suc_Diff1: "x \<in># M \<Longrightarrow> Suc (size (M - {#x#})) = size M"
 | |
| 4723 | using arg_cong[OF insert_DiffM, of _ _ size] by simp | |
| 4724 | ||
| 4725 | lemma size_Diff_singleton: "x \<in># M \<Longrightarrow> size (M - {#x#}) = size M - 1"
 | |
| 68406 | 4726 | by (simp flip: size_Suc_Diff1) | 
| 65547 | 4727 | |
| 4728 | lemma size_Diff_singleton_if: "size (A - {#x#}) = (if x \<in># A then size A - 1 else size A)"
 | |
| 4729 | by (simp add: diff_single_trivial size_Diff_singleton) | |
| 4730 | ||
| 4731 | lemma size_Un_Int: "size A + size B = size (A \<union># B) + size (A \<inter># B)" | |
| 4732 | by (metis inter_subset_eq_union size_union subset_mset.diff_add union_diff_inter_eq_sup) | |
| 4733 | ||
| 4734 | lemma size_Un_disjoint: "A \<inter># B = {#} \<Longrightarrow> size (A \<union># B) = size A + size B"
 | |
| 4735 | using size_Un_Int[of A B] by simp | |
| 4736 | ||
| 4737 | lemma size_Diff_subset_Int: "size (M - M') = size M - size (M \<inter># M')" | |
| 4738 | by (metis diff_intersect_left_idem size_Diff_submset subset_mset.inf_le1) | |
| 4739 | ||
| 4740 | lemma diff_size_le_size_Diff: "size (M :: _ multiset) - size M' \<le> size (M - M')" | |
| 4741 | by (simp add: diff_le_mono2 size_Diff_subset_Int size_mset_mono) | |
| 4742 | ||
| 4743 | lemma size_Diff1_less: "x\<in># M \<Longrightarrow> size (M - {#x#}) < size M"
 | |
| 4744 | by (rule Suc_less_SucD) (simp add: size_Suc_Diff1) | |
| 4745 | ||
| 4746 | lemma size_Diff2_less: "x\<in># M \<Longrightarrow> y\<in># M \<Longrightarrow> size (M - {#x#} - {#y#}) < size M"
 | |
| 4747 | by (metis less_imp_diff_less size_Diff1_less size_Diff_subset_Int) | |
| 4748 | ||
| 4749 | lemma size_Diff1_le: "size (M - {#x#}) \<le> size M"
 | |
| 4750 | by (cases "x \<in># M") (simp_all add: size_Diff1_less less_imp_le diff_single_trivial) | |
| 4751 | ||
| 4752 | lemma size_psubset: "M \<subseteq># M' \<Longrightarrow> size M < size M' \<Longrightarrow> M \<subset># M'" | |
| 4753 | using less_irrefl subset_mset_def by blast | |
| 4754 | ||
| 76700 
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
 blanchet parents: 
76682diff
changeset | 4755 | lifting_update multiset.lifting | 
| 
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
 blanchet parents: 
76682diff
changeset | 4756 | lifting_forget multiset.lifting | 
| 
c48fe2be847f
added lifting_forget as suggested by Peter Lammich
 blanchet parents: 
76682diff
changeset | 4757 | |
| 56656 | 4758 | hide_const (open) wcount | 
| 4759 | ||
| 55129 
26bd1cba3ab5
killed 'More_BNFs' by moving its various bits where they (now) belong
 blanchet parents: 
54868diff
changeset | 4760 | end |