diff -r 6cd6c553b480 -r 82a36e3d1b55 src/HOL/Relation.thy --- a/src/HOL/Relation.thy Tue Dec 06 20:08:51 2022 +0100 +++ b/src/HOL/Relation.thy Tue Dec 06 18:56:28 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"