# HG changeset patch # User desharna # Date 1669035184 -3600 # Node ID 3fc92362fbb51862a60ac13b1b7259675b7e25a6 # Parent 15f868460de94e04898e1f4429ec64dcdbca5495 strengthened and renamed lemma reflp_on_equality diff -r 15f868460de9 -r 3fc92362fbb5 NEWS --- a/NEWS Mon Nov 21 13:48:58 2022 +0100 +++ b/NEWS Mon Nov 21 13:53:04 2022 +0100 @@ -26,8 +26,9 @@ Minor INCOMPATIBILITY. * Theory "HOL.Relation": - - Strengthened lemmas. Minor INCOMPATIBILITY. + - Strengthened (and renamed) lemmas. Minor INCOMPATIBILITY. total_on_singleton + reflp_equality[simp] ~> reflp_on_equality[simp] - Added lemmas. antisym_if_asymp antisymp_if_asymp diff -r 15f868460de9 -r 3fc92362fbb5 src/HOL/Relation.thy --- a/src/HOL/Relation.thy Mon Nov 21 13:48:58 2022 +0100 +++ b/src/HOL/Relation.thy Mon Nov 21 13:53:04 2022 +0100 @@ -246,8 +246,8 @@ "refl_on A r \ (\(x, y) \ r. x \ A \ y \ A) \ (\x \ A. (x, x) \ r)" by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2) -lemma reflp_equality [simp]: "reflp (=)" - by (simp add: reflp_def) +lemma reflp_on_equality [simp]: "reflp_on A (=)" + by (simp add: reflp_on_def) lemma reflp_on_mono: "reflp_on A R \ (\x y. x \ A \ y \ A \ R x y \ Q x y) \ reflp_on A Q"