diff -r 9137085647ee -r 52e753197496 src/HOL/Library/Multiset_Order.thy --- a/src/HOL/Library/Multiset_Order.thy Thu Apr 13 09:53:37 2023 +0200 +++ b/src/HOL/Library/Multiset_Order.thy Thu Apr 13 14:54:03 2023 +0200 @@ -489,6 +489,47 @@ by (meson less_eq_multiset_def mset_lt_single_right_iff) +subsubsection \Simplifications\ + +lemma multp\<^sub>H\<^sub>O_repeat_mset_repeat_mset[simp]: + assumes "n \ 0" + shows "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B) \ multp\<^sub>H\<^sub>O R A B" +proof (rule iffI) + assume hyp: "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B)" + hence + 1: "repeat_mset n A \ repeat_mset n B" and + 2: "\y. n * count B y < n * count A y \ (\x. R y x \ n * count A x < n * count B x)" + by (simp_all add: multp\<^sub>H\<^sub>O_def) + + from 1 \n \ 0\ have "A \ B" + by auto + + moreover from 2 \n \ 0\ have "\y. count B y < count A y \ (\x. R y x \ count A x < count B x)" + by auto + + ultimately show "multp\<^sub>H\<^sub>O R A B" + by (simp add: multp\<^sub>H\<^sub>O_def) +next + assume "multp\<^sub>H\<^sub>O R A B" + hence 1: "A \ B" and 2: "\y. count B y < count A y \ (\x. R y x \ count A x < count B x)" + by (simp_all add: multp\<^sub>H\<^sub>O_def) + + from 1 have "repeat_mset n A \ repeat_mset n B" + by (simp add: assms repeat_mset_cancel1) + + moreover from 2 have "\y. n * count B y < n * count A y \ + (\x. R y x \ n * count A x < n * count B x)" + by auto + + ultimately show "multp\<^sub>H\<^sub>O R (repeat_mset n A) (repeat_mset n B)" + by (simp add: multp\<^sub>H\<^sub>O_def) +qed + +lemma multp\<^sub>H\<^sub>O_double_double[simp]: "multp\<^sub>H\<^sub>O R (A + A) (B + B) \ multp\<^sub>H\<^sub>O R A B" + using multp\<^sub>H\<^sub>O_repeat_mset_repeat_mset[of 2] + by (simp add: numeral_Bit0) + + subsection \Simprocs\ lemma mset_le_add_iff1: