diff -r c9bba9dba73b -r f91766530e13 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Tue Sep 20 14:51:58 2016 +0200 +++ b/src/HOL/Library/Multiset.thy Sun Sep 18 17:57:55 2016 +0200 @@ -247,6 +247,14 @@ by (auto simp del: count_greater_eq_Suc_zero_iff simp: count_greater_eq_Suc_zero_iff[symmetric] split: if_splits) +lemma multiset_nonemptyE [elim]: + assumes "A \ {#}" + obtains x where "x \# A" +proof - + have "\x. x \# A" by (rule ccontr) (insert assms, auto) + with that show ?thesis by blast +qed + subsubsection \Union\ @@ -2041,6 +2049,10 @@ lemma repeat_mset_cancel2: "repeat_mset a A = repeat_mset b A \ a = b \ A = {#}" by (auto simp: multiset_eq_iff) +lemma image_replicate_mset [simp]: + "image_mset f (replicate_mset n a) = replicate_mset n (f a)" + by (induct n) simp_all + subsection \Big operators\ @@ -2209,14 +2221,20 @@ translations "\i \# A. b" \ "CONST prod_mset (CONST image_mset (\i. b) A)" -lemma (in comm_semiring_1) dvd_prod_mset: +lemma (in comm_monoid_mult) prod_mset_subset_imp_dvd: + assumes "A \# B" + shows "prod_mset A dvd prod_mset B" +proof - + from assms have "B = (B - A) + A" by (simp add: subset_mset.diff_add) + also have "prod_mset \ = prod_mset (B - A) * prod_mset A" by simp + also have "prod_mset A dvd \" by simp + finally show ?thesis . +qed + +lemma (in comm_monoid_mult) dvd_prod_mset: assumes "x \# A" shows "x dvd prod_mset A" -proof - - from assms have "A = add_mset x (A - {#x#})" by simp - then obtain B where "A = add_mset x B" .. - then show ?thesis by simp -qed + using assms prod_mset_subset_imp_dvd [of "{#x#}" A] by simp lemma (in semidom) prod_mset_zero_iff [iff]: "prod_mset A = 0 \ 0 \# A" @@ -2236,10 +2254,22 @@ shows "prod_mset (A - {#a#}) = prod_mset A div a" using assms prod_mset_diff [of "{#a#}" A] by auto +lemma (in algebraic_semidom) is_unit_prod_mset_iff: + "is_unit (prod_mset A) \ (\x \# A. is_unit x)" + by (induct A) (auto simp: is_unit_mult_iff) + +lemma (in normalization_semidom) normalize_prod_mset: + "normalize (prod_mset A) = prod_mset (image_mset normalize A)" + by (induct A) (simp_all add: normalize_mult) + lemma (in normalization_semidom) normalized_prod_msetI: assumes "\a. a \# A \ normalize a = a" shows "normalize (prod_mset A) = prod_mset A" - using assms by (induct A) (simp_all add: normalize_mult) +proof - + from assms have "image_mset normalize A = A" + by (induct A) simp_all + then show ?thesis by (simp add: normalize_prod_mset) +qed subsection \Alternative representations\