--- a/NEWS Sat Jun 04 17:48:58 2022 +0200
+++ b/NEWS Sat Jun 04 17:58:08 2022 +0200
@@ -52,6 +52,8 @@
preorder.asymp_less
reflp_onD
reflp_onI
+ reflp_on_Inf
+ reflp_on_Sup
reflp_on_inf
reflp_on_mono
reflp_on_subset
--- a/src/HOL/Relation.thy Sat Jun 04 17:48:58 2022 +0200
+++ b/src/HOL/Relation.thy Sat Jun 04 17:58:08 2022 +0200
@@ -224,9 +224,15 @@
lemma refl_on_INTER: "\<forall>x\<in>S. refl_on (A x) (r x) \<Longrightarrow> refl_on (\<Inter>(A ` S)) (\<Inter>(r ` S))"
unfolding refl_on_def by fast
+lemma reflp_on_Inf: "\<forall>x\<in>S. reflp_on (A x) (R x) \<Longrightarrow> reflp_on (\<Inter>(A ` S)) (\<Sqinter>(R ` S))"
+ by (auto intro: reflp_onI dest: reflp_onD)
+
lemma refl_on_UNION: "\<forall>x\<in>S. refl_on (A x) (r x) \<Longrightarrow> refl_on (\<Union>(A ` S)) (\<Union>(r ` S))"
unfolding refl_on_def by blast
+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 {} {}"
by (simp add: refl_on_def)