src/HOL/Relation.thy
changeset 52730 6bf02eb4ddf7
parent 52392 ee996ca08de3
child 52749 ed416f4ac34e
--- a/src/HOL/Relation.thy	Thu Jul 25 08:57:16 2013 +0200
+++ b/src/HOL/Relation.thy	Thu Jul 25 12:25:07 2013 +0200
@@ -131,7 +131,6 @@
 lemma SUP_Sup_eq2 [pred_set_conv]: "(\<Squnion>i\<in>S. (\<lambda>x y. (x, y) \<in> i)) = (\<lambda>x y. (x, y) \<in> \<Union>S)"
   by (simp add: fun_eq_iff)
 
-
 subsection {* Properties of relations *}
 
 subsubsection {* Reflexivity *}
@@ -706,6 +705,12 @@
 lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)"
   by blast
 
+lemma converse_mono: "r^-1 \<subseteq> s ^-1 \<longleftrightarrow> r \<subseteq> s"
+  by auto
+
+lemma conversep_mono: "r^--1 \<le> s ^--1 \<longleftrightarrow> r \<le> s"
+  by (fact converse_mono[to_pred])
+
 lemma converse_Id [simp]: "Id^-1 = Id"
   by blast