# HG changeset patch # User desharna # Date 1741971854 -3600 # Node ID f7e08143eea2bee00b4b7e04ae23aed19819c38e # Parent 4355a2afa7a3132d18b3d6057b95c87b0f488ec9 strengthened lemma refl_on_empty[simp] diff -r 4355a2afa7a3 -r f7e08143eea2 NEWS --- a/NEWS Fri Mar 14 18:02:16 2025 +0100 +++ b/NEWS Fri Mar 14 18:04:14 2025 +0100 @@ -20,6 +20,8 @@ - Changed definition of predicate refl_on to not constrain the domain and range of the relation. To get the old semantics, explicitely use the formula "r \ A \ A \ refl_on A r". INCOMPATIBILITY. + - Strengthened lemmas. Minor INCOMPATIBILITY. + refl_on_empty[simp] - Added lemmas. reflp_on_refl_on_eq[pred_set_conv] diff -r 4355a2afa7a3 -r f7e08143eea2 src/HOL/Relation.thy --- a/src/HOL/Relation.thy Fri Mar 14 18:02:16 2025 +0100 +++ b/src/HOL/Relation.thy Fri Mar 14 18:04:14 2025 +0100 @@ -239,7 +239,7 @@ lemma reflp_on_Sup: "\x\S. reflp_on (A x) (R x) \ reflp_on (\(A ` S)) (\(R ` S))" by (auto intro: reflp_onI dest: reflp_onD) -lemma refl_on_empty [simp]: "refl_on {} {}" +lemma refl_on_empty [simp]: "refl_on {} r" by (simp add: refl_on_def) lemma reflp_on_empty [simp]: "reflp_on {} R"