# HG changeset patch # User Andreas Lochbihler # Date 1469797251 -7200 # Node ID eca71be9c94874ab9abfef65a39de4bfa204a8b7 # Parent 3e3097ac37d1e601bc25dd1674586402a4b6df39# Parent 0bcd79da075b0e2ffe65a1a47fee095cedc6ff62 merged diff -r 0bcd79da075b -r eca71be9c948 NEWS --- a/NEWS Fri Jul 29 12:44:22 2016 +0200 +++ b/NEWS Fri Jul 29 15:00:51 2016 +0200 @@ -433,8 +433,12 @@ linordered_cancel_ab_semigroup_add, and ordered_ab_semigroup_monoid_add_imp_le. INCOMPATIBILITY. -* There are some new simplification rules about multisets and the multiset -ordering. +* There are some new simplification rules about multisets, the multiset +ordering, and the subset ordering on multisets. +INCOMPATIBILITY. + +* The subset ordering on multisets has now the interpretation +ordered_ab_semigroup_monoid_add_imp_le. INCOMPATIBILITY. * Compound constants INFIMUM and SUPREMUM are mere abbreviations now. diff -r 0bcd79da075b -r eca71be9c948 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Fri Jul 29 12:44:22 2016 +0200 +++ b/src/HOL/Library/Multiset.thy Fri Jul 29 15:00:51 2016 +0200 @@ -469,9 +469,12 @@ apply (auto intro: multiset_eq_iff [THEN iffD2]) done -interpretation subset_mset: ordered_cancel_comm_monoid_diff "op +" 0 "op \#" "op <#" "op -" +interpretation subset_mset: ordered_cancel_comm_monoid_diff "op +" 0 "op \#" "op <#" "op -" by standard (simp, fact mset_subset_eq_exists_conv) +interpretation subset_mset: ordered_ab_semigroup_monoid_add_imp_le "op +" 0 "op -" "op \#" "op <#" + by standard + lemma mset_subset_eq_mono_add_right_cancel [simp]: "(A::'a multiset) + C \# B + C \ A \# B" by (fact subset_mset.add_le_cancel_right) @@ -481,11 +484,11 @@ lemma mset_subset_eq_mono_add: "(A::'a multiset) \# B \ C \# D \ A + C \# B + D" by (fact subset_mset.add_mono) -lemma mset_subset_eq_add_left [simp]: "(A::'a multiset) \# A + B" - unfolding subseteq_mset_def by auto +lemma mset_subset_eq_add_left: "(A::'a multiset) \# A + B" + by simp -lemma mset_subset_eq_add_right [simp]: "B \# (A::'a multiset) + B" - unfolding subseteq_mset_def by auto +lemma mset_subset_eq_add_right: "B \# (A::'a multiset) + B" + by simp lemma single_subset_iff [simp]: "{#a#} \# M \ a \# M"