strengthened lemma refl_on_empty[simp]
authordesharna
Fri, 14 Mar 2025 18:04:14 +0100
changeset 82271 f7e08143eea2
parent 82270 4355a2afa7a3
child 82272 a317d9e27a03
strengthened lemma refl_on_empty[simp]
NEWS
src/HOL/Relation.thy
--- 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 \<subset> A \<times> A \<and> refl_on A r". INCOMPATIBILITY.
+  - Strengthened lemmas. Minor INCOMPATIBILITY.
+      refl_on_empty[simp]
   - Added lemmas.
       reflp_on_refl_on_eq[pred_set_conv]
 
--- 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: "\<forall>x\<in>S. reflp_on (A x) (R x) \<Longrightarrow> reflp_on (\<Union>(A ` S)) (\<Squnion>(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"