# HG changeset patch # User kleing # Date 1114682569 -7200 # Node ID 3aca7f05cd122631f202fd0da087c33a0b0c6abc # Parent 9634b3f9d9102062a310497bd1a4c58378156377 intersection diff -r 9634b3f9d910 -r 3aca7f05cd12 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Thu Apr 28 09:21:35 2005 +0200 +++ b/src/HOL/Library/Multiset.thy Thu Apr 28 12:02:49 2005 +0200 @@ -52,6 +52,10 @@ Zero_multiset_def [simp]: "0 == {#}" size_def: "size M == setsum (count M) (set_of M)" +constdefs + multiset_inter :: "'a multiset \ 'a multiset \ 'a multiset" (infixl "#\" 70) + "multiset_inter A B \ A - (A - B)" + text {* \medskip Preservation of the representing set @{term multiset}. @@ -264,6 +268,37 @@ Better: use wf_finite_psubset in WF_Rel *) +declare Rep_multiset_inject [symmetric, simp del] + + +subsubsection {* Intersection *} + +lemma multiset_inter_count: + "count (A #\ B) x = min (count A x) (count B x)" + by (simp add:multiset_inter_def min_def) + +lemma multiset_inter_commute: "A #\ B = B #\ A" + by (simp add: multiset_eq_conv_count_eq multiset_inter_count + min_max.below_inf.inf_commute) + +lemma multiset_inter_assoc: "A #\ (B #\ C) = A #\ B #\ C" + by (simp add: multiset_eq_conv_count_eq multiset_inter_count + min_max.below_inf.inf_assoc) + +lemma multiset_inter_left_commute: "A #\ (B #\ C) = B #\ (A #\ C)" + by (simp add: multiset_eq_conv_count_eq multiset_inter_count min_def) + +lemmas multiset_inter_ac = multiset_inter_commute multiset_inter_assoc + multiset_inter_left_commute + +lemma multiset_union_diff_commute: "B #\ C = {#} \ A + B - C = A - C + B" + apply (simp add:multiset_eq_conv_count_eq multiset_inter_count min_def + split:split_if_asm) + apply clarsimp + apply (erule_tac x="a" in allE) + apply auto + done + subsection {* Induction over multisets *} @@ -332,7 +367,7 @@ prefer 2 apply (simp add: expand_fun_eq) apply (erule ssubst) - apply (erule Abs_multiset_inverse [THEN subst]) + apply (erule Abs_multiset_inverse [THEN subst]) apply (erule prem2 [simplified]) done qed @@ -354,14 +389,12 @@ theorem multiset_partition: "M = {# x:M. P x #} + {# x:M. \ P x #}" by (subst multiset_eq_conv_count_eq, auto) -declare Rep_multiset_inject [symmetric, simp del] -declare multiset_typedef [simp del] - theorem add_eq_conv_ex: "(M + {#a#} = N + {#b#}) = (M = N \ a = b \ (\K. M = K + {#b#} \ N = K + {#a#}))" by (auto simp add: add_eq_conv_diff) +declare multiset_typedef [simp del] subsection {* Multiset orderings *}