# HG changeset patch # User desharna # Date 1638123312 -3600 # Node ID 7b0a241732c1a12365962501d28fe403fa656c04 # Parent 2741ef11ccf6be06af888795d9f737b0991b2ca8 added definitions multp{DM,HO} and corresponding lemmas diff -r 2741ef11ccf6 -r 7b0a241732c1 src/HOL/Library/Multiset_Order.thy --- a/src/HOL/Library/Multiset_Order.thy Sun Nov 28 19:12:48 2021 +0100 +++ b/src/HOL/Library/Multiset_Order.thy Sun Nov 28 19:15:12 2021 +0100 @@ -11,6 +11,135 @@ subsection \Alternative Characterizations\ +subsubsection \The Dershowitz--Manna Ordering\ + +definition multp\<^sub>D\<^sub>M where + "multp\<^sub>D\<^sub>M r M N \ + (\X Y. X \ {#} \ X \# N \ M = (N - X) + Y \ (\k. k \# Y \ (\a. a \# X \ r k a)))" + +lemma multp\<^sub>D\<^sub>M_imp_multp: + "multp\<^sub>D\<^sub>M r M N \ multp r M N" +proof - + assume "multp\<^sub>D\<^sub>M r M N" + then obtain X Y where + "X \ {#}" and "X \# N" and "M = N - X + Y" and "\k. k \# Y \ (\a. a \# X \ r k a)" + unfolding multp\<^sub>D\<^sub>M_def by blast + then have "multp r (N - X + Y) (N - X + X)" + by (intro one_step_implies_multp) (auto simp: Bex_def trans_def) + with \M = N - X + Y\ \X \# N\ show "multp r M N" + by (metis subset_mset.diff_add) +qed + +subsubsection \The Huet--Oppen Ordering\ + +definition multp\<^sub>H\<^sub>O where + "multp\<^sub>H\<^sub>O r M N \ M \ N \ (\y. count N y < count M y \ (\x. r y x \ count M x < count N x))" + +lemma multp_imp_multp\<^sub>H\<^sub>O: + assumes "asymp r" and "transp r" + shows "multp r M N \ multp\<^sub>H\<^sub>O r M N" + unfolding multp_def mult_def +proof (induction rule: trancl_induct) + case (base P) + then show ?case + using \asymp r\ + by (auto elim!: mult1_lessE simp: count_eq_zero_iff multp\<^sub>H\<^sub>O_def split: if_splits + dest!: Suc_lessD) +next + case (step N P) + from step(3) have "M \ N" and + **: "\y. count N y < count M y \ (\x. r y x \ count M x < count N x)" + by (simp_all add: multp\<^sub>H\<^sub>O_def) + from step(2) obtain M0 a K where + *: "P = add_mset a M0" "N = M0 + K" "a \# K" "\b. b \# K \ r b a" + using \asymp r\ by (auto elim: mult1_lessE) + from \M \ N\ ** *(1,2,3) have "M \ P" + using *(4) \asymp r\ + by (metis asymp.cases add_cancel_right_right add_diff_cancel_left' add_mset_add_single count_inI + count_union diff_diff_add_mset diff_single_trivial in_diff_count multi_member_last) + moreover + { assume "count P a \ count M a" + with \a \# K\ have "count N a < count M a" unfolding *(1,2) + by (auto simp add: not_in_iff) + with ** obtain z where z: "r a z" "count M z < count N z" + by blast + with * have "count N z \ count P z" + using \asymp r\ + by (metis add_diff_cancel_left' add_mset_add_single asymp.cases diff_diff_add_mset + diff_single_trivial in_diff_count not_le_imp_less) + with z have "\z. r a z \ count M z < count P z" by auto + } note count_a = this + { fix y + assume count_y: "count P y < count M y" + have "\x. r y x \ count M x < count P x" + proof (cases "y = a") + case True + with count_y count_a show ?thesis by auto + next + case False + show ?thesis + proof (cases "y \# K") + case True + with *(4) have "r y a" by simp + then show ?thesis + by (cases "count P a \ count M a") (auto dest: count_a intro: \transp r\[THEN transpD]) + next + case False + with \y \ a\ have "count P y = count N y" unfolding *(1,2) + by (simp add: not_in_iff) + with count_y ** obtain z where z: "r y z" "count M z < count N z" by auto + show ?thesis + proof (cases "z \# K") + case True + with *(4) have "r z a" by simp + with z(1) show ?thesis + by (cases "count P a \ count M a") (auto dest!: count_a intro: \transp r\[THEN transpD]) + next + case False + with \a \# K\ have "count N z \ count P z" unfolding * + by (auto simp add: not_in_iff) + with z show ?thesis by auto + qed + qed + qed + } + ultimately show ?case unfolding multp\<^sub>H\<^sub>O_def by blast +qed + +lemma multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M: "multp\<^sub>H\<^sub>O r M N \ multp\<^sub>D\<^sub>M r M N" +unfolding multp\<^sub>D\<^sub>M_def +proof (intro iffI exI conjI) + assume "multp\<^sub>H\<^sub>O r M N" + then obtain z where z: "count M z < count N z" + unfolding multp\<^sub>H\<^sub>O_def by (auto simp: multiset_eq_iff nat_neq_iff) + define X where "X = N - M" + define Y where "Y = M - N" + from z show "X \ {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq) + from z show "X \# N" unfolding X_def by auto + show "M = (N - X) + Y" unfolding X_def Y_def multiset_eq_iff count_union count_diff by force + show "\k. k \# Y \ (\a. a \# X \ r k a)" + proof (intro allI impI) + fix k + assume "k \# Y" + then have "count N k < count M k" unfolding Y_def + by (auto simp add: in_diff_count) + with \multp\<^sub>H\<^sub>O r M N\ obtain a where "r k a" and "count M a < count N a" + unfolding multp\<^sub>H\<^sub>O_def by blast + then show "\a. a \# X \ r k a" unfolding X_def + by (auto simp add: in_diff_count) + qed +qed + +lemma multp_eq_multp\<^sub>D\<^sub>M: "asymp r \ transp r \ multp r = multp\<^sub>D\<^sub>M r" + using multp\<^sub>D\<^sub>M_imp_multp multp_imp_multp\<^sub>H\<^sub>O[THEN multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M] + by blast + +lemma multp_eq_multp\<^sub>H\<^sub>O: "asymp r \ transp r \ multp r = multp\<^sub>H\<^sub>O r" + using multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M[THEN multp\<^sub>D\<^sub>M_imp_multp] multp_imp_multp\<^sub>H\<^sub>O + by blast + +subsubsection \Properties of Preorders\ + context preorder begin @@ -59,107 +188,29 @@ lemma mult_imp_less_multiset\<^sub>H\<^sub>O: "(M, N) \ mult {(x, y). x < y} \ less_multiset\<^sub>H\<^sub>O M N" -proof (unfold mult_def, induct rule: trancl_induct) - case (base P) - then show ?case - by (auto elim!: mult1_lessE simp add: count_eq_zero_iff less_multiset\<^sub>H\<^sub>O_def split: if_splits dest!: Suc_lessD) -next - case (step N P) - from step(3) have "M \ N" and - **: "\y. count N y < count M y \ (\x>y. count M x < count N x)" - by (simp_all add: less_multiset\<^sub>H\<^sub>O_def) - from step(2) obtain M0 a K where - *: "P = add_mset a M0" "N = M0 + K" "a \# K" "\b. b \# K \ b < a" - by (auto elim: mult1_lessE) - from \M \ N\ ** *(1,2,3) have "M \ P" by (force dest: *(4) elim!: less_asym split: if_splits ) - moreover - { assume "count P a \ count M a" - with \a \# K\ have "count N a < count M a" unfolding *(1,2) - by (auto simp add: not_in_iff) - with ** obtain z where z: "z > a" "count M z < count N z" - by blast - with * have "count N z \ count P z" - by (auto elim: less_asym intro: count_inI) - with z have "\z > a. count M z < count P z" by auto - } note count_a = this - { fix y - assume count_y: "count P y < count M y" - have "\x>y. count M x < count P x" - proof (cases "y = a") - case True - with count_y count_a show ?thesis by auto - next - case False - show ?thesis - proof (cases "y \# K") - case True - with *(4) have "y < a" by simp - then show ?thesis by (cases "count P a \ count M a") (auto dest: count_a intro: less_trans) - next - case False - with \y \ a\ have "count P y = count N y" unfolding *(1,2) - by (simp add: not_in_iff) - with count_y ** obtain z where z: "z > y" "count M z < count N z" by auto - show ?thesis - proof (cases "z \# K") - case True - with *(4) have "z < a" by simp - with z(1) show ?thesis - by (cases "count P a \ count M a") (auto dest!: count_a intro: less_trans) - next - case False - with \a \# K\ have "count N z \ count P z" unfolding * - by (auto simp add: not_in_iff) - with z show ?thesis by auto - qed - qed - qed - } - ultimately show ?case unfolding less_multiset\<^sub>H\<^sub>O_def by blast -qed + unfolding multp_def[of "(<)", symmetric] + using multp_imp_multp\<^sub>H\<^sub>O[of "(<)"] + by (simp add: less_multiset\<^sub>H\<^sub>O_def multp\<^sub>H\<^sub>O_def) lemma less_multiset\<^sub>D\<^sub>M_imp_mult: "less_multiset\<^sub>D\<^sub>M M N \ (M, N) \ mult {(x, y). x < y}" -proof - - assume "less_multiset\<^sub>D\<^sub>M M N" - then obtain X Y where - "X \ {#}" and "X \# N" and "M = N - X + Y" and "\k. k \# Y \ (\a. a \# X \ k < a)" - unfolding less_multiset\<^sub>D\<^sub>M_def by blast - then have "(N - X + Y, N - X + X) \ mult {(x, y). x < y}" - by (intro one_step_implies_mult) (auto simp: Bex_def trans_def) - with \M = N - X + Y\ \X \# N\ show "(M, N) \ mult {(x, y). x < y}" - by (metis subset_mset.diff_add) -qed + unfolding multp_def[of "(<)", symmetric] + by (rule multp\<^sub>D\<^sub>M_imp_multp[of "(<)" M N]) (simp add: less_multiset\<^sub>D\<^sub>M_def multp\<^sub>D\<^sub>M_def) lemma less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M: "less_multiset\<^sub>H\<^sub>O M N \ less_multiset\<^sub>D\<^sub>M M N" -unfolding less_multiset\<^sub>D\<^sub>M_def -proof (intro iffI exI conjI) - assume "less_multiset\<^sub>H\<^sub>O M N" - then obtain z where z: "count M z < count N z" - unfolding less_multiset\<^sub>H\<^sub>O_def by (auto simp: multiset_eq_iff nat_neq_iff) - define X where "X = N - M" - define Y where "Y = M - N" - from z show "X \ {#}" unfolding X_def by (auto simp: multiset_eq_iff not_less_eq_eq Suc_le_eq) - from z show "X \# N" unfolding X_def by auto - show "M = (N - X) + Y" unfolding X_def Y_def multiset_eq_iff count_union count_diff by force - show "\k. k \# Y \ (\a. a \# X \ k < a)" - proof (intro allI impI) - fix k - assume "k \# Y" - then have "count N k < count M k" unfolding Y_def - by (auto simp add: in_diff_count) - with \less_multiset\<^sub>H\<^sub>O M N\ obtain a where "k < a" and "count M a < count N a" - unfolding less_multiset\<^sub>H\<^sub>O_def by blast - then show "\a. a \# X \ k < a" unfolding X_def - by (auto simp add: in_diff_count) - qed -qed + unfolding less_multiset\<^sub>D\<^sub>M_def less_multiset\<^sub>H\<^sub>O_def + unfolding multp\<^sub>D\<^sub>M_def[symmetric] multp\<^sub>H\<^sub>O_def[symmetric] + by (rule multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M) lemma mult_less_multiset\<^sub>D\<^sub>M: "(M, N) \ mult {(x, y). x < y} \ less_multiset\<^sub>D\<^sub>M M N" - by (metis less_multiset\<^sub>D\<^sub>M_imp_mult less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M mult_imp_less_multiset\<^sub>H\<^sub>O) + unfolding multp_def[of "(<)", symmetric] + using multp_eq_multp\<^sub>D\<^sub>M[of "(<)", simplified] + by (simp add: multp\<^sub>D\<^sub>M_def less_multiset\<^sub>D\<^sub>M_def) lemma mult_less_multiset\<^sub>H\<^sub>O: "(M, N) \ mult {(x, y). x < y} \ less_multiset\<^sub>H\<^sub>O M N" - by (metis less_multiset\<^sub>D\<^sub>M_imp_mult less_multiset\<^sub>H\<^sub>O_imp_less_multiset\<^sub>D\<^sub>M mult_imp_less_multiset\<^sub>H\<^sub>O) + unfolding multp_def[of "(<)", symmetric] + using multp_eq_multp\<^sub>H\<^sub>O[of "(<)", simplified] + by (simp add: multp\<^sub>H\<^sub>O_def less_multiset\<^sub>H\<^sub>O_def) lemmas mult\<^sub>D\<^sub>M = mult_less_multiset\<^sub>D\<^sub>M[unfolded less_multiset\<^sub>D\<^sub>M_def] lemmas mult\<^sub>H\<^sub>O = mult_less_multiset\<^sub>H\<^sub>O[unfolded less_multiset\<^sub>H\<^sub>O_def]