changeset 82271 | f7e08143eea2 |
parent 82270 | 4355a2afa7a3 |
child 82272 | a317d9e27a03 |
--- 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"