theory Multiset_new imports "~~/src/HOL/Library/Multiset" begin subsection \The multiset extension is cancellative for multiset union\ lemma mult_cancel: assumes "trans s" and "irrefl s" shows "(X + Z, Y + Z) \ mult s \ (X, Y) \ mult s" (is "?L \ ?R") proof assume ?L thus ?R proof (induct Z) case (add Z z) obtain X' Y' Z' where *: "X + Z + {#z#} = Z' + X'" "Y + Z + {#z#} = Z' + Y'" "Y' \ {#}" "\x \ set_mset X'. \y \ set_mset Y'. (x, y) \ s" using mult_implies_one_step[OF `trans s` add(2)] unfolding union_assoc by blast consider Z2 where "Z' = Z2 + {#z#}" | X2 Y2 where "X' = X2 + {#z#}" "Y' = Y2 + {#z#}" using *(1,2) by (metis mset_add union_iff union_single_eq_member) thus ?case proof (cases) case 1 thus ?thesis using * one_step_implies_mult[of Y' X' s Z2] by (auto simp: union_commute[of _ "{#_#}"] union_assoc intro: add(1)) next case 2 then obtain y where "y \ set_mset Y2" "(z, y) \ s" using *(4) `irrefl s` by (auto simp: irrefl_def) moreover with transD[OF `trans s` _ this(2)] have "x' \ set_mset X2 \ \y \ set_mset Y2. (x', y) \ s" for x' using 2 *(4)[rule_format, of x'] by auto ultimately show ?thesis using * one_step_implies_mult[of Y2 X2 s Z'] 2 by (force simp: union_commute[of "{#_#}"] union_assoc[symmetric] intro: add(1)) qed qed auto next assume ?R then obtain I J K where "Y = I + J" "X = I + K" "J \ {#}" "\k \ set_mset K. \j \ set_mset J. (k, j) \ s" using mult_implies_one_step[OF `trans s`] by blast thus ?L using one_step_implies_mult[of J K s "I + Z"] by (auto simp: ac_simps) qed lemma mult_cancel_max: assumes "trans s" and "irrefl s" shows "(X, Y) \ mult s \ (X - X #\ Y, Y - X #\ Y) \ mult s" (is "?L \ ?R") proof - have "X - X #\ Y + X #\ Y = X" "Y - X #\ Y + X #\ Y = Y" by (auto simp: count_inject[symmetric]) thus ?thesis using mult_cancel[OF assms, of "X - X #\ Y" "X #\ Y" "Y - X #\ Y"] by auto qed subsection \Quasi-executable version of the multiset extension\ text \ Predicate variants of \mult\ and the reflexive closure of \mult\, which are executable whenever the given predicate \P\ is. Together with the standard code equations for \op #\\ and \op -\ this should yield a quadratic (with respect to calls to \P\) implementation of \multeqp\. \ definition multp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp P N M = (let Z = M #\ N; X = M - Z; Y = N - Z in X \ {#} \ (\y \ set_mset Y. \x \ set_mset X. P y x))" definition multeqp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multeqp P N M = (let Z = M #\ N; X = M - Z; Y = N - Z in (\y \ set_mset Y. \x \ set_mset X. P y x))" lemma multp_iff: assumes "irrefl R" and "trans R" and [simp]: "\x y. P x y \ (x, y) \ R" shows "multp P N M \ (N, M) \ mult R" (is "?L \ ?R") proof - have *: "M #\ N + (N - M #\ N) = N" "M #\ N + (M - M #\ N) = M" "(M - M #\ N) #\ (N - M #\ N) = {#}" by (auto simp: count_inject[symmetric]) show ?thesis proof assume ?L thus ?R using one_step_implies_mult[of "M - M #\ N" "N - M #\ N" R "M #\ N"] * by (auto simp: multp_def Let_def) next { fix I J K :: "'a multiset" assume "(I + J) #\ (I + K) = {#}" then have "count I x = 0" for x by (auto simp: count_inject[symmetric] dest!: fun_cong[of _ _ x]) then have "I = {#}" by (auto simp: count_inject[symmetric]) } note [dest!] = this assume ?R thus ?L using mult_implies_one_step[OF assms(2), of "N - M #\ N" "M - M #\ N"] mult_cancel_max[OF assms(2,1), of "N" "M"] * by (auto simp: multp_def ac_simps) qed qed lemma multeqp_iff: assumes "irrefl R" and "trans R" and "\x y. P x y \ (x, y) \ R" shows "multeqp P N M \ (N, M) \ (mult R)\<^sup>=" proof - { assume "N \ M" "M - M #\ N = {#}" then obtain y where "count N y \ count M y" by (auto simp: count_inject[symmetric]) then have "\y. count M y < count N y" using `M - M #\ N = {#}` by (auto simp: count_inject[symmetric] dest!: le_neq_implies_less fun_cong[of _ _ y]) } then have "multeqp P N M \ multp P N M \ N = M" by (auto simp: multeqp_def multp_def Let_def in_diff_count) thus ?thesis using multp_iff[OF assms] by simp qed subsection \Monotonicity of the multiset extension\ lemma mono_mult1: assumes "s \ s'" shows "mult1 s \ mult1 s'" unfolding mult1_def using assms by blast lemma mono_mult: assumes "s \ s'" shows "mult s \ mult s'" unfolding mult_def using mono_mult1[OF assms] trancl_mono by blast end