# HG changeset patch # User wenzelm # Date 1670443397 -3600 # Node ID 186dcfe746e33ba6ff957dd4a273d556dd0f15b1 # Parent 1c083e32aed6380b24dd5dbe8a48ce9fcdaac73d# Parent badb5264f7b9b0969dfd282e488ddf805de18fd6 merged diff -r badb5264f7b9 -r 186dcfe746e3 src/HOL/Library/Multiset.thy --- a/src/HOL/Library/Multiset.thy Wed Dec 07 21:02:43 2022 +0100 +++ b/src/HOL/Library/Multiset.thy Wed Dec 07 21:03:17 2022 +0100 @@ -2951,6 +2951,8 @@ definition multp :: "('a \ 'a \ bool) \ 'a multiset \ 'a multiset \ bool" where "multp r M N \ (M, N) \ mult {(x, y). r x y}" +declare multp_def[pred_set_conv] + lemma mult1I: assumes "M = add_mset a M0" and "N = M0 + K" and "\b. b \# K \ (b, a) \ r" shows "(N, M) \ mult1 r" @@ -3138,8 +3140,9 @@ qed qed -lemmas multp_implies_one_step = - mult_implies_one_step[of "{(x, y). r x y}" for r, folded multp_def transp_trans, simplified] +lemma multp_implies_one_step: + "transp R \ multp R M N \ \I J K. N = I + J \ M = I + K \ J \ {#} \ (\k\#K. \x\#J. R k x)" + by (rule mult_implies_one_step[to_pred]) lemma one_step_implies_mult: assumes @@ -3173,8 +3176,9 @@ qed qed -lemmas one_step_implies_multp = - one_step_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def, simplified] +lemma one_step_implies_multp: + "J \ {#} \ \k\#K. \j\#J. R k j \ multp R (I + K) (I + J)" + by (rule one_step_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def, simplified]) lemma subset_implies_mult: assumes sub: "A \# B" @@ -3188,7 +3192,8 @@ by (rule one_step_implies_mult[of "B - A" "{#}" _ A, unfolded ApBmA, simplified]) qed -lemmas subset_implies_multp = subset_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def] +lemma subset_implies_multp: "A \# B \ multp r A B" + by (rule subset_implies_mult[of _ _ "{(x, y). r x y}" for r, folded multp_def]) subsubsection \Monotonicity\ @@ -3342,16 +3347,17 @@ thus ?L using one_step_implies_mult[of J K s "I + Z"] by (auto simp: ac_simps) qed -lemmas multp_cancel = - mult_cancel[of "{(x, y). r x y}" for r, - folded multp_def transp_trans irreflp_irrefl_eq, simplified] - -lemmas mult_cancel_add_mset = - mult_cancel[of _ _ "{#_#}", unfolded union_mset_add_mset_right add.comm_neutral] - -lemmas multp_cancel_add_mset = - mult_cancel_add_mset[of "{(x, y). r x y}" for r, - folded multp_def transp_trans irreflp_irrefl_eq, simplified] +lemma multp_cancel: + "transp R \ irreflp R \ multp R (X + Z) (Y + Z) \ multp R X Y" + by (rule mult_cancel[to_pred]) + +lemma mult_cancel_add_mset: + "trans r \ irrefl r \ ((add_mset z X, add_mset z Y) \ mult r) = ((X, Y) \ mult r)" + by (rule mult_cancel[of _ _ "{#_#}", unfolded union_mset_add_mset_right add.comm_neutral]) + +lemma multp_cancel_add_mset: + "transp R \ irreflp R \ multp R (add_mset z X) (add_mset z Y) = multp R X Y" + by (rule mult_cancel_add_mset[to_pred]) lemma mult_cancel_max0: assumes "trans s" and "irrefl s" @@ -3363,9 +3369,8 @@ lemmas mult_cancel_max = mult_cancel_max0[simplified] -lemmas multp_cancel_max = - mult_cancel_max[of "{(x, y). r x y}" for r, - folded multp_def transp_trans irreflp_irrefl_eq, simplified] +lemma multp_cancel_max: "transp R \ irreflp R \ multp R X Y = multp R (X - Y) (Y - X)" + by (rule mult_cancel_max[to_pred]) subsubsection \Partial-order properties\ @@ -3416,9 +3421,9 @@ with * show False by simp qed -lemmas irreflp_multp = - irrefl_mult[of "{(x, y). r x y}" for r, - folded transp_trans_eq irreflp_irrefl_eq, simplified, folded multp_def] +lemma irreflp_multp: "transp R \ irreflp R \ irreflp (multp R)" + by (rule irrefl_mult[of "{(x, y). r x y}" for r, + folded transp_trans_eq irreflp_irrefl_eq, simplified, folded multp_def]) instantiation multiset :: (preorder) order begin diff -r badb5264f7b9 -r 186dcfe746e3 src/HOL/Relation.thy --- a/src/HOL/Relation.thy Wed Dec 07 21:02:43 2022 +0100 +++ b/src/HOL/Relation.thy Wed Dec 07 21:03:17 2022 +0100 @@ -298,10 +298,10 @@ by (rule irrefl_onI[of UNIV, simplified]) lemma irreflp_onI: "(\a. a \ A \ \ R a a) \ irreflp_on A R" - by (simp add: irreflp_on_def) + by (rule irrefl_onI[to_pred]) lemma irreflpI[intro?]: "(\a. \ R a a) \ irreflp R" - by (rule irreflp_onI[of UNIV, simplified]) + by (rule irreflI[to_pred]) lemma irrefl_onD: "irrefl_on A r \ a \ A \ (a, a) \ r" by (simp add: irrefl_on_def) @@ -310,10 +310,10 @@ by (rule irrefl_onD[of UNIV, simplified]) lemma irreflp_onD: "irreflp_on A R \ a \ A \ \ R a a" - by (simp add: irreflp_on_def) + by (rule irrefl_onD[to_pred]) lemma irreflpD: "irreflp R \ \ R x x" - by (rule irreflp_onD[of UNIV, simplified]) + by (rule irreflD[to_pred]) lemma irrefl_on_distinct [code]: "irrefl_on A r \ (\(a, b) \ r. a \ A \ b \ A \ a \ b)" by (auto simp add: irrefl_on_def) @@ -609,10 +609,10 @@ by (rule total_onI) lemma totalp_onI: "(\x y. x \ A \ y \ A \ x \ y \ R x y \ R y x) \ totalp_on A R" - by (simp add: totalp_on_def) + by (rule total_onI[to_pred]) lemma totalpI: "(\x y. x \ y \ R x y \ R y x) \ totalp R" - by (rule totalp_onI) + by (rule totalI[to_pred]) lemma totalp_onD: "totalp_on A R \ x \ A \ y \ A \ x \ y \ R x y \ R y x" diff -r badb5264f7b9 -r 186dcfe746e3 src/HOL/Wellfounded.thy --- a/src/HOL/Wellfounded.thy Wed Dec 07 21:02:43 2022 +0100 +++ b/src/HOL/Wellfounded.thy Wed Dec 07 21:03:17 2022 +0100 @@ -75,7 +75,7 @@ using wf_irrefl [OF assms] by (auto simp add: irrefl_def) lemma wfP_imp_irreflp: "wfP r \ irreflp r" - by (rule wf_imp_irrefl[to_pred, folded top_set_def]) + by (rule wf_imp_irrefl[to_pred]) lemma wf_wellorderI: assumes wf: "wf {(x::'a::ord, y). x < y}"