diff -r abe0c3872d8a -r 76302202a92d src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Wed Aug 10 18:57:20 2016 +0200 +++ b/src/HOL/Library/Multiset.thy Mon Aug 08 14:01:49 2016 +0200 @@ -2214,6 +2214,14 @@ obtains a M0 K where "M = M0 + {#a#}" "N = M0 + K" "\b. b \# K \ (b, a) \ r" using assms unfolding mult1_def by blast +lemma mono_mult1: + assumes "r \ r'" shows "mult1 r \ mult1 r'" +unfolding mult1_def using assms by blast + +lemma mono_mult: + assumes "r \ r'" shows "mult r \ mult r'" +unfolding mult_def using mono_mult1[OF assms] trancl_mono by blast + lemma not_less_empty [iff]: "(M, {#}) \ mult1 r" by (simp add: mult1_def) @@ -2396,115 +2404,103 @@ \ (I + K, I + J) \ mult r" using one_step_implies_mult_aux by blast -subsection \A quasi-executable characterization\ + +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 add.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: add.commute[of _ "{#_#}"] add.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 from this 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: add.commute[of "{#_#}"] add.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 \ - The decreasing parts \A\ and \B\ of multisets in a multiset-comparison - \(I + B, I + A) \ mult r\, can always be made disjoint. + 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 quadratic + (with respect to calls to \P\) implementations of \multp\ and \multeqp\. \ -lemma decreasing_parts_disj: - assumes "irrefl r" and "trans r" - and "A \ {#}" and *: "\b\#B. \a\#A. (b, a) \ r" - defines "Z \ A #\ B" - defines "X \ A - Z" - defines "Y \ B - Z" - shows "X \ {#} \ X #\ Y = {#} \ - A = X + Z \ B = Y + Z \ (\y\#Y. \x\#X. (y, x) \ r)" -proof - - define D - where "D = set_mset A \ set_mset B" - let ?r = "r \ D \ D" - have "irrefl ?r" and "trans ?r" and "finite ?r" - using \irrefl r\ and \trans r\ by (auto simp: D_def irrefl_def trans_Restr) - note wf_converse_induct = wf_induct [OF wf_converse [OF this]] - { fix b assume "b \# B" - then have "\x. x \# X \ (b, x) \ r" - proof (induction rule: wf_converse_induct) - case (1 b) - then obtain a where "b \# B" and "a \# A" and "(b, a) \ r" - using * by blast - then show ?case - proof (cases "a \# X") - case False - then have "a \# B" using \a \# A\ - by (simp add: X_def Z_def not_in_iff) - (metis le_zero_eq not_in_iff) - moreover have "(a, b) \ (r \ D \ D)\" - using \(b, a) \ r\ using \b \# B\ and \a \# B\ - by (auto simp: D_def) - ultimately obtain x where x: "x \# X" "(a, x) \ r" - using "1.IH" by blast - then have "(b, x) \ r" - using \trans r\ and \(b, a) \ r\ by (auto dest: transD) - with x show ?thesis by blast - qed blast - qed } - note B_less = this - then have "\y\#Y. \x\#X. (y, x) \ r" - by (auto simp: Y_def Z_def dest: in_diffD) - moreover have "X \ {#}" - proof - - obtain a where "a \# A" using \A \ {#}\ - by (auto simp: multiset_eq_iff) - show ?thesis - proof (cases "a \# X") - case False - then have "a \# B" using \a \# A\ - by (simp add: X_def Z_def not_in_iff) - (metis le_zero_eq not_in_iff) - then show ?thesis by (auto dest: B_less) - qed auto - qed - moreover have "A = X + Z" and "B = Y + Z" and "X #\ Y = {#}" - by (auto simp: X_def Y_def Z_def multiset_eq_iff) - ultimately show ?thesis by blast -qed - -text \ - A predicate variant of the reflexive closure of \mult\, which is - 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 in + X \ {#} \ (let Y = N - Z in (\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 multeqp_iff: - assumes "irrefl R" and "trans R" - and [simp]: "\x y. P x y \ (x, y) \ R" - shows "multeqp P N M \ (N, M) \ (mult R)\<^sup>=" -proof - assume "(N, M) \ (mult R)\<^sup>=" - then show "multeqp P N M" +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 "(N, M) \ mult R" - from mult_implies_one_step [OF \trans R\ this] obtain I J K - where *: "I \ {#}" "\j\#J. \i\#I. (j, i) \ R" - and [simp]: "M = K + I" "N = K + J" by blast - from decreasing_parts_disj [OF \irrefl R\ \trans R\ *] - show "multeqp P N M" - by (auto simp: multeqp_def split: if_splits) + 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 - assume "(N, M) \ Id" then show "multeqp P N M" by (auto simp: multeqp_def) + { fix I J K :: "'a multiset" assume "(I + J) #\ (I + K) = {#}" + then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty) + } 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 -next - assume "multeqp P N M" - then obtain X Y Z where *: "Z = M #\ N" "X = M - Z" "Y = N - Z" - and **: "\y\#Y. \x\#X. (y, x) \ R" by (auto simp: multeqp_def Let_def) - then have M: "M = Z + X" and N: "N = Z + Y" by (auto simp: multiset_eq_iff) - show "(N, M) \ (mult R)\<^sup>=" - proof (cases "X \ {#}") - case True - with * and ** have "(Z + Y, Z + X) \ mult R" - by (auto intro: one_step_implies_mult) - then show ?thesis by (simp add: M N) - next - case False - then show ?thesis using ** - by (cases "Y = {#}") (auto simp: M N) - 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