diff -r 4dd01eb1e0cb -r d10a49b7b620 src/HOL/Relation.thy --- a/src/HOL/Relation.thy Mon Mar 17 09:00:40 2025 +0100 +++ b/src/HOL/Relation.thy Mon Mar 17 09:12:18 2025 +0100 @@ -203,6 +203,12 @@ obtains "r x x" using assms by (auto dest: refl_onD simp add: reflp_def) +lemma refl_on_top[simp]: "refl_on A \" + by (simp add: refl_on_def) + +lemma reflp_on_top[simp]: "reflp_on A \" + by (simp add: reflp_on_def) + lemma reflp_on_subset: "reflp_on A R \ B \ A \ reflp_on B R" by (auto intro: reflp_onI dest: reflp_onD) @@ -452,6 +458,12 @@ lemma symp_on_bot[simp]: "symp_on A \" using sym_on_bot[to_pred] . +lemma sym_on_top[simp]: "sym_on A \" + by (simp add: sym_on_def) + +lemma symp_on_top[simp]: "symp_on A \" + by (simp add: symp_on_def) + lemma sym_on_subset: "sym_on A r \ B \ A \ sym_on B r" by (auto simp: sym_on_def) @@ -742,6 +754,12 @@ lemma transp_on_bot[simp]: "transp_on A \" using trans_on_bot[to_pred] . +lemma trans_on_top[simp]: "trans_on A \" + by (simp add: trans_on_def) + +lemma transp_on_top[simp]: "transp_on A \" + by (simp add: transp_on_def) + lemma transp_empty [simp]: "transp (\x y. False)" using transp_on_bot unfolding bot_fun_def bot_bool_def . @@ -807,6 +825,12 @@ lemma totalpD: "totalp R \ x \ y \ R x y \ R y x" by (simp add: totalp_onD) +lemma total_on_top[simp]: "total_on A \" + by (simp add: total_on_def) + +lemma totalp_on_top[simp]: "totalp_on A \" + by (simp add: totalp_on_def) + lemma totalp_on_mono_stronger: fixes A :: "'a set" and R :: "'a \ 'a \ bool" and