src/HOL/Relation.thy
changeset 63404 a95e7432d86c
parent 63376 4c0cc2b356f0
child 63561 fba08009ff3e
--- a/src/HOL/Relation.thy	Wed Jul 06 14:09:13 2016 +0200
+++ b/src/HOL/Relation.thy	Wed Jul 06 20:19:51 2016 +0200
@@ -14,7 +14,7 @@
 declare predicate1D [Pure.dest, dest]
 declare predicate2I [Pure.intro!, intro!]
 declare predicate2D [Pure.dest, dest]
-declare bot1E [elim!] 
+declare bot1E [elim!]
 declare bot2E [elim!]
 declare top1I [intro!]
 declare top2I [intro!]
@@ -56,15 +56,16 @@
 
 subsubsection \<open>Relations as sets of pairs\<close>
 
-type_synonym 'a rel = "('a * 'a) set"
+type_synonym 'a rel = "('a \<times> 'a) set"
 
-lemma subrelI: \<comment> \<open>Version of @{thm [source] subsetI} for binary relations\<close>
-  "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
+lemma subrelI: "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (x, y) \<in> s) \<Longrightarrow> r \<subseteq> s"
+  \<comment> \<open>Version of @{thm [source] subsetI} for binary relations\<close>
   by auto
 
-lemma lfp_induct2: \<comment> \<open>Version of @{thm [source] lfp_induct} for binary relations\<close>
+lemma lfp_induct2:
   "(a, b) \<in> lfp f \<Longrightarrow> mono f \<Longrightarrow>
     (\<And>a b. (a, b) \<in> f (lfp f \<inter> {(x, y). P x y}) \<Longrightarrow> P a b) \<Longrightarrow> P a b"
+  \<comment> \<open>Version of @{thm [source] lfp_induct} for binary relations\<close>
   using lfp_induct_set [of "(a, b)" f "case_prod P"] by auto
 
 
@@ -148,35 +149,30 @@
 subsubsection \<open>Reflexivity\<close>
 
 definition refl_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
-where
-  "refl_on A r \<longleftrightarrow> r \<subseteq> A \<times> A \<and> (\<forall>x\<in>A. (x, x) \<in> r)"
+  where "refl_on A r \<longleftrightarrow> r \<subseteq> A \<times> A \<and> (\<forall>x\<in>A. (x, x) \<in> r)"
 
-abbreviation refl :: "'a rel \<Rightarrow> bool"
-where \<comment> \<open>reflexivity over a type\<close>
-  "refl \<equiv> refl_on UNIV"
+abbreviation refl :: "'a rel \<Rightarrow> bool" \<comment> \<open>reflexivity over a type\<close>
+  where "refl \<equiv> refl_on UNIV"
 
 definition reflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where
-  "reflp r \<longleftrightarrow> (\<forall>x. r x x)"
+  where "reflp r \<longleftrightarrow> (\<forall>x. r x x)"
 
-lemma reflp_refl_eq [pred_set_conv]:
-  "reflp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> refl r" 
+lemma reflp_refl_eq [pred_set_conv]: "reflp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> refl r"
   by (simp add: refl_on_def reflp_def)
 
-lemma refl_onI [intro?]: "r \<subseteq> A \<times> A ==> (!!x. x : A ==> (x, x) : r) ==> refl_on A r"
-  by (unfold refl_on_def) (iprover intro!: ballI)
+lemma refl_onI [intro?]: "r \<subseteq> A \<times> A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> (x, x) \<in> r) \<Longrightarrow> refl_on A r"
+  unfolding refl_on_def by (iprover intro!: ballI)
 
-lemma refl_onD: "refl_on A r ==> a : A ==> (a, a) : r"
-  by (unfold refl_on_def) blast
+lemma refl_onD: "refl_on A r \<Longrightarrow> a \<in> A \<Longrightarrow> (a, a) \<in> r"
+  unfolding refl_on_def by blast
 
-lemma refl_onD1: "refl_on A r ==> (x, y) : r ==> x : A"
-  by (unfold refl_on_def) blast
+lemma refl_onD1: "refl_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> x \<in> A"
+  unfolding refl_on_def by blast
 
-lemma refl_onD2: "refl_on A r ==> (x, y) : r ==> y : A"
-  by (unfold refl_on_def) blast
+lemma refl_onD2: "refl_on A r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> y \<in> A"
+  unfolding refl_on_def by blast
 
-lemma reflpI [intro?]:
-  "(\<And>x. r x x) \<Longrightarrow> reflp r"
+lemma reflpI [intro?]: "(\<And>x. r x x) \<Longrightarrow> reflp r"
   by (auto intro: refl_onI simp add: reflp_def)
 
 lemma reflpE:
@@ -189,104 +185,86 @@
   shows "r x x"
   using assms by (auto elim: reflpE)
 
-lemma refl_on_Int: "refl_on A r ==> refl_on B s ==> refl_on (A \<inter> B) (r \<inter> s)"
-  by (unfold refl_on_def) blast
+lemma refl_on_Int: "refl_on A r \<Longrightarrow> refl_on B s \<Longrightarrow> refl_on (A \<inter> B) (r \<inter> s)"
+  unfolding refl_on_def by blast
 
-lemma reflp_inf:
-  "reflp r \<Longrightarrow> reflp s \<Longrightarrow> reflp (r \<sqinter> s)"
+lemma reflp_inf: "reflp r \<Longrightarrow> reflp s \<Longrightarrow> reflp (r \<sqinter> s)"
   by (auto intro: reflpI elim: reflpE)
 
-lemma refl_on_Un: "refl_on A r ==> refl_on B s ==> refl_on (A \<union> B) (r \<union> s)"
-  by (unfold refl_on_def) blast
+lemma refl_on_Un: "refl_on A r \<Longrightarrow> refl_on B s \<Longrightarrow> refl_on (A \<union> B) (r \<union> s)"
+  unfolding refl_on_def by blast
 
-lemma reflp_sup:
-  "reflp r \<Longrightarrow> reflp s \<Longrightarrow> reflp (r \<squnion> s)"
+lemma reflp_sup: "reflp r \<Longrightarrow> reflp s \<Longrightarrow> reflp (r \<squnion> s)"
   by (auto intro: reflpI elim: reflpE)
 
-lemma refl_on_INTER:
-  "ALL x:S. refl_on (A x) (r x) ==> refl_on (INTER S A) (INTER S r)"
-  by (unfold refl_on_def) fast
+lemma refl_on_INTER: "\<forall>x\<in>S. refl_on (A x) (r x) \<Longrightarrow> refl_on (INTER S A) (INTER S r)"
+  unfolding refl_on_def by fast
 
-lemma refl_on_UNION:
-  "ALL x:S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
-  by (unfold refl_on_def) blast
+lemma refl_on_UNION: "\<forall>x\<in>S. refl_on (A x) (r x) \<Longrightarrow> refl_on (UNION S A) (UNION S r)"
+  unfolding refl_on_def by blast
 
 lemma refl_on_empty [simp]: "refl_on {} {}"
-  by (simp add:refl_on_def)
+  by (simp add: refl_on_def)
 
 lemma refl_on_def' [nitpick_unfold, code]:
   "refl_on A r \<longleftrightarrow> (\<forall>(x, y) \<in> r. x \<in> A \<and> y \<in> A) \<and> (\<forall>x \<in> A. (x, x) \<in> r)"
   by (auto intro: refl_onI dest: refl_onD refl_onD1 refl_onD2)
 
 lemma reflp_equality [simp]: "reflp op ="
-by(simp add: reflp_def)
+  by (simp add: reflp_def)
 
-lemma reflp_mono: "\<lbrakk> reflp R; \<And>x y. R x y \<longrightarrow> Q x y \<rbrakk> \<Longrightarrow> reflp Q"
-by(auto intro: reflpI dest: reflpD)
+lemma reflp_mono: "reflp R \<Longrightarrow> (\<And>x y. R x y \<longrightarrow> Q x y) \<Longrightarrow> reflp Q"
+  by (auto intro: reflpI dest: reflpD)
 
 
 subsubsection \<open>Irreflexivity\<close>
 
 definition irrefl :: "'a rel \<Rightarrow> bool"
-where
-  "irrefl r \<longleftrightarrow> (\<forall>a. (a, a) \<notin> r)"
+  where "irrefl r \<longleftrightarrow> (\<forall>a. (a, a) \<notin> r)"
 
 definition irreflp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where
-  "irreflp R \<longleftrightarrow> (\<forall>a. \<not> R a a)"
+  where "irreflp R \<longleftrightarrow> (\<forall>a. \<not> R a a)"
 
-lemma irreflp_irrefl_eq [pred_set_conv]:
-  "irreflp (\<lambda>a b. (a, b) \<in> R) \<longleftrightarrow> irrefl R" 
+lemma irreflp_irrefl_eq [pred_set_conv]: "irreflp (\<lambda>a b. (a, b) \<in> R) \<longleftrightarrow> irrefl R"
   by (simp add: irrefl_def irreflp_def)
 
-lemma irreflI [intro?]:
-  "(\<And>a. (a, a) \<notin> R) \<Longrightarrow> irrefl R"
+lemma irreflI [intro?]: "(\<And>a. (a, a) \<notin> R) \<Longrightarrow> irrefl R"
   by (simp add: irrefl_def)
 
-lemma irreflpI [intro?]:
-  "(\<And>a. \<not> R a a) \<Longrightarrow> irreflp R"
+lemma irreflpI [intro?]: "(\<And>a. \<not> R a a) \<Longrightarrow> irreflp R"
   by (fact irreflI [to_pred])
 
-lemma irrefl_distinct [code]:
-  "irrefl r \<longleftrightarrow> (\<forall>(a, b) \<in> r. a \<noteq> b)"
+lemma irrefl_distinct [code]: "irrefl r \<longleftrightarrow> (\<forall>(a, b) \<in> r. a \<noteq> b)"
   by (auto simp add: irrefl_def)
 
 
 subsubsection \<open>Asymmetry\<close>
 
 inductive asym :: "'a rel \<Rightarrow> bool"
-where
-  asymI: "irrefl R \<Longrightarrow> (\<And>a b. (a, b) \<in> R \<Longrightarrow> (b, a) \<notin> R) \<Longrightarrow> asym R"
+  where asymI: "irrefl R \<Longrightarrow> (\<And>a b. (a, b) \<in> R \<Longrightarrow> (b, a) \<notin> R) \<Longrightarrow> asym R"
 
 inductive asymp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where
-  asympI: "irreflp R \<Longrightarrow> (\<And>a b. R a b \<Longrightarrow> \<not> R b a) \<Longrightarrow> asymp R"
+  where asympI: "irreflp R \<Longrightarrow> (\<And>a b. R a b \<Longrightarrow> \<not> R b a) \<Longrightarrow> asymp R"
 
-lemma asymp_asym_eq [pred_set_conv]:
-  "asymp (\<lambda>a b. (a, b) \<in> R) \<longleftrightarrow> asym R" 
+lemma asymp_asym_eq [pred_set_conv]: "asymp (\<lambda>a b. (a, b) \<in> R) \<longleftrightarrow> asym R"
   by (auto intro!: asymI asympI elim: asym.cases asymp.cases simp add: irreflp_irrefl_eq)
 
 
 subsubsection \<open>Symmetry\<close>
 
 definition sym :: "'a rel \<Rightarrow> bool"
-where
-  "sym r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r)"
+  where "sym r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r)"
 
 definition symp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where
-  "symp r \<longleftrightarrow> (\<forall>x y. r x y \<longrightarrow> r y x)"
+  where "symp r \<longleftrightarrow> (\<forall>x y. r x y \<longrightarrow> r y x)"
 
-lemma symp_sym_eq [pred_set_conv]:
-  "symp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> sym r" 
+lemma symp_sym_eq [pred_set_conv]: "symp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> sym r"
   by (simp add: sym_def symp_def)
 
-lemma symI [intro?]:
-  "(\<And>a b. (a, b) \<in> r \<Longrightarrow> (b, a) \<in> r) \<Longrightarrow> sym r"
+lemma symI [intro?]: "(\<And>a b. (a, b) \<in> r \<Longrightarrow> (b, a) \<in> r) \<Longrightarrow> sym r"
   by (unfold sym_def) iprover
 
-lemma sympI [intro?]:
-  "(\<And>a b. r a b \<Longrightarrow> r b a) \<Longrightarrow> symp r"
+lemma sympI [intro?]: "(\<And>a b. r a b \<Longrightarrow> r b a) \<Longrightarrow> symp r"
   by (fact symI [to_pred])
 
 lemma symE:
@@ -309,86 +287,70 @@
   shows "r a b"
   using assms by (rule symD [to_pred])
 
-lemma sym_Int:
-  "sym r \<Longrightarrow> sym s \<Longrightarrow> sym (r \<inter> s)"
+lemma sym_Int: "sym r \<Longrightarrow> sym s \<Longrightarrow> sym (r \<inter> s)"
   by (fast intro: symI elim: symE)
 
-lemma symp_inf:
-  "symp r \<Longrightarrow> symp s \<Longrightarrow> symp (r \<sqinter> s)"
+lemma symp_inf: "symp r \<Longrightarrow> symp s \<Longrightarrow> symp (r \<sqinter> s)"
   by (fact sym_Int [to_pred])
 
-lemma sym_Un:
-  "sym r \<Longrightarrow> sym s \<Longrightarrow> sym (r \<union> s)"
+lemma sym_Un: "sym r \<Longrightarrow> sym s \<Longrightarrow> sym (r \<union> s)"
   by (fast intro: symI elim: symE)
 
-lemma symp_sup:
-  "symp r \<Longrightarrow> symp s \<Longrightarrow> symp (r \<squnion> s)"
+lemma symp_sup: "symp r \<Longrightarrow> symp s \<Longrightarrow> symp (r \<squnion> s)"
   by (fact sym_Un [to_pred])
 
-lemma sym_INTER:
-  "\<forall>x\<in>S. sym (r x) \<Longrightarrow> sym (INTER S r)"
+lemma sym_INTER: "\<forall>x\<in>S. sym (r x) \<Longrightarrow> sym (INTER S r)"
   by (fast intro: symI elim: symE)
 
-lemma symp_INF:
-  "\<forall>x\<in>S. symp (r x) \<Longrightarrow> symp (INFIMUM S r)"
+lemma symp_INF: "\<forall>x\<in>S. symp (r x) \<Longrightarrow> symp (INFIMUM S r)"
   by (fact sym_INTER [to_pred])
 
-lemma sym_UNION:
-  "\<forall>x\<in>S. sym (r x) \<Longrightarrow> sym (UNION S r)"
+lemma sym_UNION: "\<forall>x\<in>S. sym (r x) \<Longrightarrow> sym (UNION S r)"
   by (fast intro: symI elim: symE)
 
-lemma symp_SUP:
-  "\<forall>x\<in>S. symp (r x) \<Longrightarrow> symp (SUPREMUM S r)"
+lemma symp_SUP: "\<forall>x\<in>S. symp (r x) \<Longrightarrow> symp (SUPREMUM S r)"
   by (fact sym_UNION [to_pred])
 
 
 subsubsection \<open>Antisymmetry\<close>
 
 definition antisym :: "'a rel \<Rightarrow> bool"
-where
-  "antisym r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r \<longrightarrow> x = y)"
+  where "antisym r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (y, x) \<in> r \<longrightarrow> x = y)"
 
 abbreviation antisymP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where -- \<open>FIXME proper logical operation\<close>
-  "antisymP r \<equiv> antisym {(x, y). r x y}"
+  where "antisymP r \<equiv> antisym {(x, y). r x y}" (* FIXME proper logical operation *)
+
+lemma antisymI [intro?]: "(\<And>x y. (x, y) \<in> r \<Longrightarrow> (y, x) \<in> r \<Longrightarrow> x = y) \<Longrightarrow> antisym r"
+  unfolding antisym_def by iprover
 
-lemma antisymI [intro?]:
-  "(!!x y. (x, y) : r ==> (y, x) : r ==> x=y) ==> antisym r"
-  by (unfold antisym_def) iprover
+lemma antisymD [dest?]: "antisym r \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> (b, a) \<in> r \<Longrightarrow> a = b"
+  unfolding antisym_def by iprover
 
-lemma antisymD [dest?]: "antisym r ==> (a, b) : r ==> (b, a) : r ==> a = b"
-  by (unfold antisym_def) iprover
-
-lemma antisym_subset: "r \<subseteq> s ==> antisym s ==> antisym r"
-  by (unfold antisym_def) blast
+lemma antisym_subset: "r \<subseteq> s \<Longrightarrow> antisym s \<Longrightarrow> antisym r"
+  unfolding antisym_def by blast
 
 lemma antisym_empty [simp]: "antisym {}"
-  by (unfold antisym_def) blast
+  unfolding antisym_def by blast
 
 lemma antisymP_equality [simp]: "antisymP op ="
-by(auto intro: antisymI)
+  by (auto intro: antisymI)
 
 
 subsubsection \<open>Transitivity\<close>
 
 definition trans :: "'a rel \<Rightarrow> bool"
-where
-  "trans r \<longleftrightarrow> (\<forall>x y z. (x, y) \<in> r \<longrightarrow> (y, z) \<in> r \<longrightarrow> (x, z) \<in> r)"
+  where "trans r \<longleftrightarrow> (\<forall>x y z. (x, y) \<in> r \<longrightarrow> (y, z) \<in> r \<longrightarrow> (x, z) \<in> r)"
 
 definition transp :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where
-  "transp r \<longleftrightarrow> (\<forall>x y z. r x y \<longrightarrow> r y z \<longrightarrow> r x z)"
+  where "transp r \<longleftrightarrow> (\<forall>x y z. r x y \<longrightarrow> r y z \<longrightarrow> r x z)"
 
-lemma transp_trans_eq [pred_set_conv]:
-  "transp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> trans r" 
+lemma transp_trans_eq [pred_set_conv]: "transp (\<lambda>x y. (x, y) \<in> r) \<longleftrightarrow> trans r"
   by (simp add: trans_def transp_def)
 
-lemma transI [intro?]:
-  "(\<And>x y z. (x, y) \<in> r \<Longrightarrow> (y, z) \<in> r \<Longrightarrow> (x, z) \<in> r) \<Longrightarrow> trans r"
+lemma transI [intro?]: "(\<And>x y z. (x, y) \<in> r \<Longrightarrow> (y, z) \<in> r \<Longrightarrow> (x, z) \<in> r) \<Longrightarrow> trans r"
   by (unfold trans_def) iprover
 
-lemma transpI [intro?]:
-  "(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r"
+lemma transpI [intro?]: "(\<And>x y z. r x y \<Longrightarrow> r y z \<Longrightarrow> r x z) \<Longrightarrow> transp r"
   by (fact transI [to_pred])
 
 lemma transE:
@@ -411,37 +373,31 @@
   shows "r x z"
   using assms by (rule transD [to_pred])
 
-lemma trans_Int:
-  "trans r \<Longrightarrow> trans s \<Longrightarrow> trans (r \<inter> s)"
+lemma trans_Int: "trans r \<Longrightarrow> trans s \<Longrightarrow> trans (r \<inter> s)"
   by (fast intro: transI elim: transE)
 
-lemma transp_inf:
-  "transp r \<Longrightarrow> transp s \<Longrightarrow> transp (r \<sqinter> s)"
+lemma transp_inf: "transp r \<Longrightarrow> transp s \<Longrightarrow> transp (r \<sqinter> s)"
   by (fact trans_Int [to_pred])
 
-lemma trans_INTER:
-  "\<forall>x\<in>S. trans (r x) \<Longrightarrow> trans (INTER S r)"
+lemma trans_INTER: "\<forall>x\<in>S. trans (r x) \<Longrightarrow> trans (INTER S r)"
   by (fast intro: transI elim: transD)
 
 (* FIXME thm trans_INTER [to_pred] *)
 
-lemma trans_join [code]:
-  "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
+lemma trans_join [code]: "trans r \<longleftrightarrow> (\<forall>(x, y1) \<in> r. \<forall>(y2, z) \<in> r. y1 = y2 \<longrightarrow> (x, z) \<in> r)"
   by (auto simp add: trans_def)
 
-lemma transp_trans:
-  "transp r \<longleftrightarrow> trans {(x, y). r x y}"
+lemma transp_trans: "transp r \<longleftrightarrow> trans {(x, y). r x y}"
   by (simp add: trans_def transp_def)
 
 lemma transp_equality [simp]: "transp op ="
-by(auto intro: transpI)
+  by (auto intro: transpI)
 
 
 subsubsection \<open>Totality\<close>
 
 definition total_on :: "'a set \<Rightarrow> 'a rel \<Rightarrow> bool"
-where
-  "total_on A r \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. x \<noteq> y \<longrightarrow> (x, y) \<in> r \<or> (y, x) \<in> r)"
+  where "total_on A r \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. x \<noteq> y \<longrightarrow> (x, y) \<in> r \<or> (y, x) \<in> r)"
 
 abbreviation "total \<equiv> total_on UNIV"
 
@@ -452,27 +408,22 @@
 subsubsection \<open>Single valued relations\<close>
 
 definition single_valued :: "('a \<times> 'b) set \<Rightarrow> bool"
-where
-  "single_valued r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (\<forall>z. (x, z) \<in> r \<longrightarrow> y = z))"
+  where "single_valued r \<longleftrightarrow> (\<forall>x y. (x, y) \<in> r \<longrightarrow> (\<forall>z. (x, z) \<in> r \<longrightarrow> y = z))"
 
 abbreviation single_valuedP :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> bool"
-where -- \<open>FIXME proper logical operation\<close>
-  "single_valuedP r \<equiv> single_valued {(x, y). r x y}"
+  where "single_valuedP r \<equiv> single_valued {(x, y). r x y}" (* FIXME proper logical operation *)
 
-lemma single_valuedI:
-  "ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z) ==> single_valued r"
-  by (unfold single_valued_def)
+lemma single_valuedI: "\<forall>x y. (x, y) \<in> r \<longrightarrow> (\<forall>z. (x, z) \<in> r \<longrightarrow> y = z) \<Longrightarrow> single_valued r"
+  unfolding single_valued_def .
 
-lemma single_valuedD:
-  "single_valued r ==> (x, y) : r ==> (x, z) : r ==> y = z"
+lemma single_valuedD: "single_valued r \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> (x, z) \<in> r \<Longrightarrow> y = z"
   by (simp add: single_valued_def)
 
 lemma single_valued_empty[simp]: "single_valued {}"
-by(simp add: single_valued_def)
+  by (simp add: single_valued_def)
 
-lemma single_valued_subset:
-  "r \<subseteq> s ==> single_valued s ==> single_valued r"
-  by (unfold single_valued_def) blast
+lemma single_valued_subset: "r \<subseteq> s \<Longrightarrow> single_valued s \<Longrightarrow> single_valued r"
+  unfolding single_valued_def by blast
 
 
 subsection \<open>Relation operations\<close>
@@ -480,17 +431,16 @@
 subsubsection \<open>The identity relation\<close>
 
 definition Id :: "'a rel"
-where
-  [code del]: "Id = {p. \<exists>x. p = (x, x)}"
+  where [code del]: "Id = {p. \<exists>x. p = (x, x)}"
 
-lemma IdI [intro]: "(a, a) : Id"
+lemma IdI [intro]: "(a, a) \<in> Id"
   by (simp add: Id_def)
 
-lemma IdE [elim!]: "p : Id ==> (!!x. p = (x, x) ==> P) ==> P"
-  by (unfold Id_def) (iprover elim: CollectE)
+lemma IdE [elim!]: "p \<in> Id \<Longrightarrow> (\<And>x. p = (x, x) \<Longrightarrow> P) \<Longrightarrow> P"
+  unfolding Id_def by (iprover elim: CollectE)
 
-lemma pair_in_Id_conv [iff]: "((a, b) : Id) = (a = b)"
-  by (unfold Id_def) blast
+lemma pair_in_Id_conv [iff]: "(a, b) \<in> Id \<longleftrightarrow> a = b"
+  unfolding Id_def by blast
 
 lemma refl_Id: "refl Id"
   by (simp add: refl_on_def)
@@ -509,7 +459,7 @@
   by (unfold single_valued_def) blast
 
 lemma irrefl_diff_Id [simp]: "irrefl (r - Id)"
-  by (simp add:irrefl_def)
+  by (simp add: irrefl_def)
 
 lemma trans_diff_Id: "trans r \<Longrightarrow> antisym r \<Longrightarrow> trans (r - Id)"
   unfolding antisym_def trans_def by blast
@@ -524,28 +474,25 @@
 subsubsection \<open>Diagonal: identity over a set\<close>
 
 definition Id_on  :: "'a set \<Rightarrow> 'a rel"
-where
-  "Id_on A = (\<Union>x\<in>A. {(x, x)})"
+  where "Id_on A = (\<Union>x\<in>A. {(x, x)})"
 
 lemma Id_on_empty [simp]: "Id_on {} = {}"
-  by (simp add: Id_on_def) 
+  by (simp add: Id_on_def)
 
-lemma Id_on_eqI: "a = b ==> a : A ==> (a, b) : Id_on A"
+lemma Id_on_eqI: "a = b \<Longrightarrow> a \<in> A \<Longrightarrow> (a, b) \<in> Id_on A"
   by (simp add: Id_on_def)
 
-lemma Id_onI [intro!]: "a : A ==> (a, a) : Id_on A"
+lemma Id_onI [intro!]: "a \<in> A \<Longrightarrow> (a, a) \<in> Id_on A"
   by (rule Id_on_eqI) (rule refl)
 
-lemma Id_onE [elim!]:
-  "c : Id_on A ==> (!!x. x : A ==> c = (x, x) ==> P) ==> P"
+lemma Id_onE [elim!]: "c \<in> Id_on A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> c = (x, x) \<Longrightarrow> P) \<Longrightarrow> P"
   \<comment> \<open>The general elimination rule.\<close>
-  by (unfold Id_on_def) (iprover elim!: UN_E singletonE)
+  unfolding Id_on_def by (iprover elim!: UN_E singletonE)
 
-lemma Id_on_iff: "((x, y) : Id_on A) = (x = y & x : A)"
+lemma Id_on_iff: "(x, y) \<in> Id_on A \<longleftrightarrow> x = y \<and> x \<in> A"
   by blast
 
-lemma Id_on_def' [nitpick_unfold]:
-  "Id_on {x. A x} = Collect (\<lambda>(x, y). x = y \<and> A x)"
+lemma Id_on_def' [nitpick_unfold]: "Id_on {x. A x} = Collect (\<lambda>(x, y). x = y \<and> A x)"
   by auto
 
 lemma Id_on_subset_Times: "Id_on A \<subseteq> A \<times> A"
@@ -555,7 +502,7 @@
   by (rule refl_onI [OF Id_on_subset_Times Id_onI])
 
 lemma antisym_Id_on [simp]: "antisym (Id_on A)"
-  by (unfold antisym_def) blast
+  unfolding antisym_def by blast
 
 lemma sym_Id_on [simp]: "sym (Id_on A)"
   by (rule symI) clarify
@@ -564,15 +511,14 @@
   by (fast intro: transI elim: transD)
 
 lemma single_valued_Id_on [simp]: "single_valued (Id_on A)"
-  by (unfold single_valued_def) blast
+  unfolding single_valued_def by blast
 
 
 subsubsection \<open>Composition\<close>
 
-inductive_set relcomp  :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'c) set \<Rightarrow> ('a \<times> 'c) set" (infixr "O" 75)
+inductive_set relcomp  :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'c) set \<Rightarrow> ('a \<times> 'c) set"  (infixr "O" 75)
   for r :: "('a \<times> 'b) set" and s :: "('b \<times> 'c) set"
-where
-  relcompI [intro]: "(a, b) \<in> r \<Longrightarrow> (b, c) \<in> s \<Longrightarrow> (a, c) \<in> r O s"
+  where relcompI [intro]: "(a, b) \<in> r \<Longrightarrow> (b, c) \<in> s \<Longrightarrow> (a, c) \<in> r O s"
 
 notation relcompp (infixr "OO" 75)
 
@@ -588,269 +534,239 @@
 
 lemma relcompE [elim!]: "xz \<in> r O s \<Longrightarrow>
   (\<And>x y z. xz = (x, z) \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> (y, z) \<in> s  \<Longrightarrow> P) \<Longrightarrow> P"
-  by (cases xz) (simp, erule relcompEpair, iprover)
+  apply (cases xz)
+  apply simp
+  apply (erule relcompEpair)
+  apply iprover
+  done
 
-lemma R_O_Id [simp]:
-  "R O Id = R"
+lemma R_O_Id [simp]: "R O Id = R"
   by fast
 
-lemma Id_O_R [simp]:
-  "Id O R = R"
+lemma Id_O_R [simp]: "Id O R = R"
   by fast
 
-lemma relcomp_empty1 [simp]:
-  "{} O R = {}"
+lemma relcomp_empty1 [simp]: "{} O R = {}"
   by blast
 
-lemma relcompp_bot1 [simp]:
-  "\<bottom> OO R = \<bottom>"
+lemma relcompp_bot1 [simp]: "\<bottom> OO R = \<bottom>"
   by (fact relcomp_empty1 [to_pred])
 
-lemma relcomp_empty2 [simp]:
-  "R O {} = {}"
+lemma relcomp_empty2 [simp]: "R O {} = {}"
   by blast
 
-lemma relcompp_bot2 [simp]:
-  "R OO \<bottom> = \<bottom>"
+lemma relcompp_bot2 [simp]: "R OO \<bottom> = \<bottom>"
   by (fact relcomp_empty2 [to_pred])
 
-lemma O_assoc:
-  "(R O S) O T = R O (S O T)"
+lemma O_assoc: "(R O S) O T = R O (S O T)"
   by blast
 
-lemma relcompp_assoc:
-  "(r OO s) OO t = r OO (s OO t)"
+lemma relcompp_assoc: "(r OO s) OO t = r OO (s OO t)"
   by (fact O_assoc [to_pred])
 
-lemma trans_O_subset:
-  "trans r \<Longrightarrow> r O r \<subseteq> r"
+lemma trans_O_subset: "trans r \<Longrightarrow> r O r \<subseteq> r"
   by (unfold trans_def) blast
 
-lemma transp_relcompp_less_eq:
-  "transp r \<Longrightarrow> r OO r \<le> r "
+lemma transp_relcompp_less_eq: "transp r \<Longrightarrow> r OO r \<le> r "
   by (fact trans_O_subset [to_pred])
 
-lemma relcomp_mono:
-  "r' \<subseteq> r \<Longrightarrow> s' \<subseteq> s \<Longrightarrow> r' O s' \<subseteq> r O s"
+lemma relcomp_mono: "r' \<subseteq> r \<Longrightarrow> s' \<subseteq> s \<Longrightarrow> r' O s' \<subseteq> r O s"
   by blast
 
-lemma relcompp_mono:
-  "r' \<le> r \<Longrightarrow> s' \<le> s \<Longrightarrow> r' OO s' \<le> r OO s "
+lemma relcompp_mono: "r' \<le> r \<Longrightarrow> s' \<le> s \<Longrightarrow> r' OO s' \<le> r OO s "
   by (fact relcomp_mono [to_pred])
 
-lemma relcomp_subset_Sigma:
-  "r \<subseteq> A \<times> B \<Longrightarrow> s \<subseteq> B \<times> C \<Longrightarrow> r O s \<subseteq> A \<times> C"
+lemma relcomp_subset_Sigma: "r \<subseteq> A \<times> B \<Longrightarrow> s \<subseteq> B \<times> C \<Longrightarrow> r O s \<subseteq> A \<times> C"
   by blast
 
-lemma relcomp_distrib [simp]:
-  "R O (S \<union> T) = (R O S) \<union> (R O T)" 
+lemma relcomp_distrib [simp]: "R O (S \<union> T) = (R O S) \<union> (R O T)"
   by auto
 
-lemma relcompp_distrib [simp]:
-  "R OO (S \<squnion> T) = R OO S \<squnion> R OO T"
+lemma relcompp_distrib [simp]: "R OO (S \<squnion> T) = R OO S \<squnion> R OO T"
   by (fact relcomp_distrib [to_pred])
 
-lemma relcomp_distrib2 [simp]:
-  "(S \<union> T) O R = (S O R) \<union> (T O R)"
+lemma relcomp_distrib2 [simp]: "(S \<union> T) O R = (S O R) \<union> (T O R)"
   by auto
 
-lemma relcompp_distrib2 [simp]:
-  "(S \<squnion> T) OO R = S OO R \<squnion> T OO R"
+lemma relcompp_distrib2 [simp]: "(S \<squnion> T) OO R = S OO R \<squnion> T OO R"
   by (fact relcomp_distrib2 [to_pred])
 
-lemma relcomp_UNION_distrib:
-  "s O UNION I r = (\<Union>i\<in>I. s O r i) "
+lemma relcomp_UNION_distrib: "s O UNION I r = (\<Union>i\<in>I. s O r i) "
   by auto
 
 (* FIXME thm relcomp_UNION_distrib [to_pred] *)
 
-lemma relcomp_UNION_distrib2:
-  "UNION I r O s = (\<Union>i\<in>I. r i O s) "
+lemma relcomp_UNION_distrib2: "UNION I r O s = (\<Union>i\<in>I. r i O s) "
   by auto
 
 (* FIXME thm relcomp_UNION_distrib2 [to_pred] *)
 
-lemma single_valued_relcomp:
-  "single_valued r \<Longrightarrow> single_valued s \<Longrightarrow> single_valued (r O s)"
-  by (unfold single_valued_def) blast
+lemma single_valued_relcomp: "single_valued r \<Longrightarrow> single_valued s \<Longrightarrow> single_valued (r O s)"
+  unfolding single_valued_def by blast
 
-lemma relcomp_unfold:
-  "r O s = {(x, z). \<exists>y. (x, y) \<in> r \<and> (y, z) \<in> s}"
+lemma relcomp_unfold: "r O s = {(x, z). \<exists>y. (x, y) \<in> r \<and> (y, z) \<in> s}"
   by (auto simp add: set_eq_iff)
 
 lemma relcompp_apply: "(R OO S) a c \<longleftrightarrow> (\<exists>b. R a b \<and> S b c)"
   unfolding relcomp_unfold [to_pred] ..
 
 lemma eq_OO: "op= OO R = R"
-by blast
+  by blast
 
 lemma OO_eq: "R OO op = = R"
-by blast
+  by blast
 
 
 subsubsection \<open>Converse\<close>
 
 inductive_set converse :: "('a \<times> 'b) set \<Rightarrow> ('b \<times> 'a) set"  ("(_\<inverse>)" [1000] 999)
   for r :: "('a \<times> 'b) set"
-where
-  "(a, b) \<in> r \<Longrightarrow> (b, a) \<in> r\<inverse>"
+  where "(a, b) \<in> r \<Longrightarrow> (b, a) \<in> r\<inverse>"
 
-notation
-  conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
+notation conversep  ("(_\<inverse>\<inverse>)" [1000] 1000)
 
 notation (ASCII)
   converse  ("(_^-1)" [1000] 999) and
   conversep ("(_^--1)" [1000] 1000)
 
-lemma converseI [sym]:
-  "(a, b) \<in> r \<Longrightarrow> (b, a) \<in> r\<inverse>"
+lemma converseI [sym]: "(a, b) \<in> r \<Longrightarrow> (b, a) \<in> r\<inverse>"
   by (fact converse.intros)
 
-lemma conversepI (* CANDIDATE [sym] *):
-  "r a b \<Longrightarrow> r\<inverse>\<inverse> b a"
+lemma conversepI (* CANDIDATE [sym] *): "r a b \<Longrightarrow> r\<inverse>\<inverse> b a"
   by (fact conversep.intros)
 
-lemma converseD [sym]:
-  "(a, b) \<in> r\<inverse> \<Longrightarrow> (b, a) \<in> r"
+lemma converseD [sym]: "(a, b) \<in> r\<inverse> \<Longrightarrow> (b, a) \<in> r"
   by (erule converse.cases) iprover
 
-lemma conversepD (* CANDIDATE [sym] *):
-  "r\<inverse>\<inverse> b a \<Longrightarrow> r a b"
+lemma conversepD (* CANDIDATE [sym] *): "r\<inverse>\<inverse> b a \<Longrightarrow> r a b"
   by (fact converseD [to_pred])
 
-lemma converseE [elim!]:
+lemma converseE [elim!]: "yx \<in> r\<inverse> \<Longrightarrow> (\<And>x y. yx = (y, x) \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> P) \<Longrightarrow> P"
   \<comment> \<open>More general than \<open>converseD\<close>, as it ``splits'' the member of the relation.\<close>
-  "yx \<in> r\<inverse> \<Longrightarrow> (\<And>x y. yx = (y, x) \<Longrightarrow> (x, y) \<in> r \<Longrightarrow> P) \<Longrightarrow> P"
-  by (cases yx) (simp, erule converse.cases, iprover)
+  apply (cases yx)
+  apply simp
+  apply (erule converse.cases)
+  apply iprover
+  done
 
 lemmas conversepE [elim!] = conversep.cases
 
-lemma converse_iff [iff]:
-  "(a, b) \<in> r\<inverse> \<longleftrightarrow> (b, a) \<in> r"
+lemma converse_iff [iff]: "(a, b) \<in> r\<inverse> \<longleftrightarrow> (b, a) \<in> r"
   by (auto intro: converseI)
 
-lemma conversep_iff [iff]:
-  "r\<inverse>\<inverse> a b = r b a"
+lemma conversep_iff [iff]: "r\<inverse>\<inverse> a b = r b a"
   by (fact converse_iff [to_pred])
 
-lemma converse_converse [simp]:
-  "(r\<inverse>)\<inverse> = r"
+lemma converse_converse [simp]: "(r\<inverse>)\<inverse> = r"
   by (simp add: set_eq_iff)
 
-lemma conversep_conversep [simp]:
-  "(r\<inverse>\<inverse>)\<inverse>\<inverse> = r"
+lemma conversep_conversep [simp]: "(r\<inverse>\<inverse>)\<inverse>\<inverse> = r"
   by (fact converse_converse [to_pred])
 
 lemma converse_empty[simp]: "{}\<inverse> = {}"
-by auto
+  by auto
 
 lemma converse_UNIV[simp]: "UNIV\<inverse> = UNIV"
-by auto
+  by auto
 
-lemma converse_relcomp: "(r O s)^-1 = s^-1 O r^-1"
+lemma converse_relcomp: "(r O s)\<inverse> = s\<inverse> O r\<inverse>"
   by blast
 
-lemma converse_relcompp: "(r OO s)^--1 = s^--1 OO r^--1"
-  by (iprover intro: order_antisym conversepI relcomppI
-    elim: relcomppE dest: conversepD)
+lemma converse_relcompp: "(r OO s)\<inverse>\<inverse> = s\<inverse>\<inverse> OO r\<inverse>\<inverse>"
+  by (iprover intro: order_antisym conversepI relcomppI elim: relcomppE dest: conversepD)
 
-lemma converse_Int: "(r \<inter> s)^-1 = r^-1 \<inter> s^-1"
+lemma converse_Int: "(r \<inter> s)\<inverse> = r\<inverse> \<inter> s\<inverse>"
   by blast
 
-lemma converse_meet: "(r \<sqinter> s)^--1 = r^--1 \<sqinter> s^--1"
+lemma converse_meet: "(r \<sqinter> s)\<inverse>\<inverse> = r\<inverse>\<inverse> \<sqinter> s\<inverse>\<inverse>"
   by (simp add: inf_fun_def) (iprover intro: conversepI ext dest: conversepD)
 
-lemma converse_Un: "(r \<union> s)^-1 = r^-1 \<union> s^-1"
+lemma converse_Un: "(r \<union> s)\<inverse> = r\<inverse> \<union> s\<inverse>"
   by blast
 
-lemma converse_join: "(r \<squnion> s)^--1 = r^--1 \<squnion> s^--1"
+lemma converse_join: "(r \<squnion> s)\<inverse>\<inverse> = r\<inverse>\<inverse> \<squnion> s\<inverse>\<inverse>"
   by (simp add: sup_fun_def) (iprover intro: conversepI ext dest: conversepD)
 
-lemma converse_INTER: "(INTER S r)^-1 = (INT x:S. (r x)^-1)"
+lemma converse_INTER: "(INTER S r)\<inverse> = (INT x:S. (r x)\<inverse>)"
   by fast
 
-lemma converse_UNION: "(UNION S r)^-1 = (UN x:S. (r x)^-1)"
+lemma converse_UNION: "(UNION S r)\<inverse> = (UN x:S. (r x)\<inverse>)"
   by blast
 
-lemma converse_mono[simp]: "r^-1 \<subseteq> s ^-1 \<longleftrightarrow> r \<subseteq> s"
+lemma converse_mono[simp]: "r\<inverse> \<subseteq> s \<inverse> \<longleftrightarrow> r \<subseteq> s"
   by auto
 
-lemma conversep_mono[simp]: "r^--1 \<le> s ^--1 \<longleftrightarrow> r \<le> s"
+lemma conversep_mono[simp]: "r\<inverse>\<inverse> \<le> s \<inverse>\<inverse> \<longleftrightarrow> r \<le> s"
   by (fact converse_mono[to_pred])
 
-lemma converse_inject[simp]: "r^-1 = s ^-1 \<longleftrightarrow> r = s"
+lemma converse_inject[simp]: "r\<inverse> = s \<inverse> \<longleftrightarrow> r = s"
   by auto
 
-lemma conversep_inject[simp]: "r^--1 = s ^--1 \<longleftrightarrow> r = s"
+lemma conversep_inject[simp]: "r\<inverse>\<inverse> = s \<inverse>\<inverse> \<longleftrightarrow> r = s"
   by (fact converse_inject[to_pred])
 
-lemma converse_subset_swap: "r \<subseteq> s ^-1 = (r ^-1 \<subseteq> s)"
+lemma converse_subset_swap: "r \<subseteq> s \<inverse> = (r \<inverse> \<subseteq> s)"
   by auto
 
-lemma conversep_le_swap: "r \<le> s ^--1 = (r ^--1 \<le> s)"
+lemma conversep_le_swap: "r \<le> s \<inverse>\<inverse> = (r \<inverse>\<inverse> \<le> s)"
   by (fact converse_subset_swap[to_pred])
 
-lemma converse_Id [simp]: "Id^-1 = Id"
+lemma converse_Id [simp]: "Id\<inverse> = Id"
   by blast
 
-lemma converse_Id_on [simp]: "(Id_on A)^-1 = Id_on A"
+lemma converse_Id_on [simp]: "(Id_on A)\<inverse> = Id_on A"
   by blast
 
 lemma refl_on_converse [simp]: "refl_on A (converse r) = refl_on A r"
-  by (unfold refl_on_def) auto
+  by (auto simp: refl_on_def)
 
 lemma sym_converse [simp]: "sym (converse r) = sym r"
-  by (unfold sym_def) blast
+  unfolding sym_def by blast
 
 lemma antisym_converse [simp]: "antisym (converse r) = antisym r"
-  by (unfold antisym_def) blast
+  unfolding antisym_def by blast
 
 lemma trans_converse [simp]: "trans (converse r) = trans r"
-  by (unfold trans_def) blast
+  unfolding trans_def by blast
 
-lemma sym_conv_converse_eq: "sym r = (r^-1 = r)"
-  by (unfold sym_def) fast
+lemma sym_conv_converse_eq: "sym r \<longleftrightarrow> r\<inverse> = r"
+  unfolding sym_def by fast
 
-lemma sym_Un_converse: "sym (r \<union> r^-1)"
-  by (unfold sym_def) blast
+lemma sym_Un_converse: "sym (r \<union> r\<inverse>)"
+  unfolding sym_def by blast
 
-lemma sym_Int_converse: "sym (r \<inter> r^-1)"
-  by (unfold sym_def) blast
+lemma sym_Int_converse: "sym (r \<inter> r\<inverse>)"
+  unfolding sym_def by blast
 
-lemma total_on_converse [simp]: "total_on A (r^-1) = total_on A r"
+lemma total_on_converse [simp]: "total_on A (r\<inverse>) = total_on A r"
   by (auto simp: total_on_def)
 
-lemma finite_converse [iff]: "finite (r^-1) = finite r"  
+lemma finite_converse [iff]: "finite (r\<inverse>) = finite r"
   unfolding converse_def conversep_iff using [[simproc add: finite_Collect]]
   by (auto elim: finite_imageD simp: inj_on_def)
 
-lemma conversep_noteq [simp]: "(op \<noteq>)^--1 = op \<noteq>"
+lemma conversep_noteq [simp]: "(op \<noteq>)\<inverse>\<inverse> = op \<noteq>"
   by (auto simp add: fun_eq_iff)
 
-lemma conversep_eq [simp]: "(op =)^--1 = op ="
+lemma conversep_eq [simp]: "(op =)\<inverse>\<inverse> = op ="
   by (auto simp add: fun_eq_iff)
 
-lemma converse_unfold [code]:
-  "r\<inverse> = {(y, x). (x, y) \<in> r}"
+lemma converse_unfold [code]: "r\<inverse> = {(y, x). (x, y) \<in> r}"
   by (simp add: set_eq_iff)
 
 
 subsubsection \<open>Domain, range and field\<close>
 
-inductive_set Domain :: "('a \<times> 'b) set \<Rightarrow> 'a set"
-  for r :: "('a \<times> 'b) set"
-where
-  DomainI [intro]: "(a, b) \<in> r \<Longrightarrow> a \<in> Domain r"
+inductive_set Domain :: "('a \<times> 'b) set \<Rightarrow> 'a set" for r :: "('a \<times> 'b) set"
+  where DomainI [intro]: "(a, b) \<in> r \<Longrightarrow> a \<in> Domain r"
 
 lemmas DomainPI = Domainp.DomainI
 
 inductive_cases DomainE [elim!]: "a \<in> Domain r"
 inductive_cases DomainpE [elim!]: "Domainp r a"
 
-inductive_set Range :: "('a \<times> 'b) set \<Rightarrow> 'b set"
-  for r :: "('a \<times> 'b) set"
-where
-  RangeI [intro]: "(a, b) \<in> r \<Longrightarrow> b \<in> Range r"
+inductive_set Range :: "('a \<times> 'b) set \<Rightarrow> 'b set" for r :: "('a \<times> 'b) set"
+  where RangeI [intro]: "(a, b) \<in> r \<Longrightarrow> b \<in> Range r"
 
 lemmas RangePI = Rangep.RangeI
 
@@ -858,15 +774,12 @@
 inductive_cases RangepE [elim!]: "Rangep r b"
 
 definition Field :: "'a rel \<Rightarrow> 'a set"
-where
-  "Field r = Domain r \<union> Range r"
+  where "Field r = Domain r \<union> Range r"
 
-lemma Domain_fst [code]:
-  "Domain r = fst ` r"
+lemma Domain_fst [code]: "Domain r = fst ` r"
   by force
 
-lemma Range_snd [code]:
-  "Range r = snd ` r"
+lemma Range_snd [code]: "Range r = snd ` r"
   by force
 
 lemma fst_eq_Domain: "fst ` R = Domain R"
@@ -962,10 +875,10 @@
 lemma Field_converse [simp]: "Field (r\<inverse>) = Field r"
   by (auto simp: Field_def)
 
-lemma Domain_Collect_case_prod [simp]: "Domain {(x, y). P x y} = {x. EX y. P x y}"
+lemma Domain_Collect_case_prod [simp]: "Domain {(x, y). P x y} = {x. \<exists>y. P x y}"
   by auto
 
-lemma Range_Collect_case_prod [simp]: "Range {(x, y). P x y} = {y. EX x. P x y}"
+lemma Range_Collect_case_prod [simp]: "Range {(x, y). P x y} = {y. \<exists>x. P x y}"
   by auto
 
 lemma finite_Domain: "finite r \<Longrightarrow> finite (Domain r)"
@@ -986,34 +899,31 @@
 lemma mono_Field: "r \<subseteq> s \<Longrightarrow> Field r \<subseteq> Field s"
   by (auto simp: Field_def Domain_def Range_def)
 
-lemma Domain_unfold:
-  "Domain r = {x. \<exists>y. (x, y) \<in> r}"
+lemma Domain_unfold: "Domain r = {x. \<exists>y. (x, y) \<in> r}"
   by blast
 
 
 subsubsection \<open>Image of a set under a relation\<close>
 
-definition Image :: "('a \<times> 'b) set \<Rightarrow> 'a set \<Rightarrow> 'b set" (infixr "``" 90)
-where
-  "r `` s = {y. \<exists>x\<in>s. (x, y) \<in> r}"
+definition Image :: "('a \<times> 'b) set \<Rightarrow> 'a set \<Rightarrow> 'b set"  (infixr "``" 90)
+  where "r `` s = {y. \<exists>x\<in>s. (x, y) \<in> r}"
 
-lemma Image_iff: "(b : r``A) = (EX x:A. (x, b) : r)"
+lemma Image_iff: "b \<in> r``A \<longleftrightarrow> (\<exists>x\<in>A. (x, b) \<in> r)"
   by (simp add: Image_def)
 
-lemma Image_singleton: "r``{a} = {b. (a, b) : r}"
+lemma Image_singleton: "r``{a} = {b. (a, b) \<in> r}"
   by (simp add: Image_def)
 
-lemma Image_singleton_iff [iff]: "(b : r``{a}) = ((a, b) : r)"
+lemma Image_singleton_iff [iff]: "b \<in> r``{a} \<longleftrightarrow> (a, b) \<in> r"
   by (rule Image_iff [THEN trans]) simp
 
-lemma ImageI [intro]: "(a, b) : r ==> a : A ==> b : r``A"
-  by (unfold Image_def) blast
+lemma ImageI [intro]: "(a, b) \<in> r \<Longrightarrow> a \<in> A \<Longrightarrow> b \<in> r``A"
+  unfolding Image_def by blast
 
-lemma ImageE [elim!]:
-  "b : r `` A ==> (!!x. (x, b) : r ==> x : A ==> P) ==> P"
-  by (unfold Image_def) (iprover elim!: CollectE bexE)
+lemma ImageE [elim!]: "b \<in> r `` A \<Longrightarrow> (\<And>x. (x, b) \<in> r \<Longrightarrow> x \<in> A \<Longrightarrow> P) \<Longrightarrow> P"
+  unfolding Image_def by (iprover elim!: CollectE bexE)
 
-lemma rev_ImageI: "a : A ==> (a, b) : r ==> b : r `` A"
+lemma rev_ImageI: "a \<in> A \<Longrightarrow> (a, b) \<in> r \<Longrightarrow> b \<in> r `` A"
   \<comment> \<open>This version's more effective when we already have the required \<open>a\<close>\<close>
   by blast
 
@@ -1029,9 +939,8 @@
 lemma Image_Int_subset: "R `` (A \<inter> B) \<subseteq> R `` A \<inter> R `` B"
   by blast
 
-lemma Image_Int_eq:
-  "single_valued (converse R) ==> R `` (A \<inter> B) = R `` A \<inter> R `` B"
-  by (simp add: single_valued_def, blast) 
+lemma Image_Int_eq: "single_valued (converse R) \<Longrightarrow> R `` (A \<inter> B) = R `` A \<inter> R `` B"
+  by (simp add: single_valued_def, blast)
 
 lemma Image_Un: "R `` (A \<union> B) = R `` A \<union> R `` B"
   by blast
@@ -1039,14 +948,14 @@
 lemma Un_Image: "(R \<union> S) `` A = R `` A \<union> S `` A"
   by blast
 
-lemma Image_subset: "r \<subseteq> A \<times> B ==> r``C \<subseteq> B"
+lemma Image_subset: "r \<subseteq> A \<times> B \<Longrightarrow> r``C \<subseteq> B"
   by (iprover intro!: subsetI elim!: ImageE dest!: subsetD SigmaD2)
 
 lemma Image_eq_UN: "r``B = (\<Union>y\<in> B. r``{y})"
   \<comment> \<open>NOT suitable for rewriting\<close>
   by blast
 
-lemma Image_mono: "r' \<subseteq> r ==> A' \<subseteq> A ==> (r' `` A') \<subseteq> (r `` A)"
+lemma Image_mono: "r' \<subseteq> r \<Longrightarrow> A' \<subseteq> A \<Longrightarrow> (r' `` A') \<subseteq> (r `` A)"
   by blast
 
 lemma Image_UN: "(r `` (UNION A B)) = (\<Union>x\<in>A. r `` (B x))"
@@ -1058,19 +967,18 @@
 lemma Image_INT_subset: "(r `` INTER A B) \<subseteq> (\<Inter>x\<in>A. r `` (B x))"
   by blast
 
-text\<open>Converse inclusion requires some assumptions\<close>
-lemma Image_INT_eq:
-  "single_valued (r\<inverse>) \<Longrightarrow> A \<noteq> {} \<Longrightarrow> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
-apply (rule equalityI)
- apply (rule Image_INT_subset) 
-apply (auto simp add: single_valued_def)
-apply blast
-done
+text \<open>Converse inclusion requires some assumptions\<close>
+lemma Image_INT_eq: "single_valued (r\<inverse>) \<Longrightarrow> A \<noteq> {} \<Longrightarrow> r `` INTER A B = (\<Inter>x\<in>A. r `` B x)"
+  apply (rule equalityI)
+   apply (rule Image_INT_subset)
+  apply (auto simp add: single_valued_def)
+  apply blast
+  done
 
-lemma Image_subset_eq: "(r``A \<subseteq> B) = (A \<subseteq> - ((r^-1) `` (-B)))"
+lemma Image_subset_eq: "r``A \<subseteq> B \<longleftrightarrow> A \<subseteq> - ((r\<inverse>) `` (- B))"
   by blast
 
-lemma Image_Collect_case_prod [simp]: "{(x, y). P x y} `` A = {y. EX x:A. P x y}"
+lemma Image_Collect_case_prod [simp]: "{(x, y). P x y} `` A = {y. \<exists>x\<in>A. P x y}"
   by auto
 
 lemma Sigma_Image: "(SIGMA x:A. B x) `` X = (\<Union>x\<in>X \<inter> A. B x)"
@@ -1083,29 +991,27 @@
 subsubsection \<open>Inverse image\<close>
 
 definition inv_image :: "'b rel \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a rel"
-where
-  "inv_image r f = {(x, y). (f x, f y) \<in> r}"
+  where "inv_image r f = {(x, y). (f x, f y) \<in> r}"
 
 definition inv_imagep :: "('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
-where
-  "inv_imagep r f = (\<lambda>x y. r (f x) (f y))"
+  where "inv_imagep r f = (\<lambda>x y. r (f x) (f y))"
 
 lemma [pred_set_conv]: "inv_imagep (\<lambda>x y. (x, y) \<in> r) f = (\<lambda>x y. (x, y) \<in> inv_image r f)"
   by (simp add: inv_image_def inv_imagep_def)
 
-lemma sym_inv_image: "sym r ==> sym (inv_image r f)"
-  by (unfold sym_def inv_image_def) blast
+lemma sym_inv_image: "sym r \<Longrightarrow> sym (inv_image r f)"
+  unfolding sym_def inv_image_def by blast
 
-lemma trans_inv_image: "trans r ==> trans (inv_image r f)"
-  apply (unfold trans_def inv_image_def)
+lemma trans_inv_image: "trans r \<Longrightarrow> trans (inv_image r f)"
+  unfolding trans_def inv_image_def
   apply (simp (no_asm))
   apply blast
   done
 
-lemma in_inv_image[simp]: "((x,y) : inv_image r f) = ((f x, f y) : r)"
+lemma in_inv_image[simp]: "(x, y) \<in> inv_image r f \<longleftrightarrow> (f x, f y) \<in> r"
   by (auto simp:inv_image_def)
 
-lemma converse_inv_image[simp]: "(inv_image R f)^-1 = inv_image (R^-1) f"
+lemma converse_inv_image[simp]: "(inv_image R f)\<inverse> = inv_image (R\<inverse>) f"
   unfolding inv_image_def converse_unfold by auto
 
 lemma in_inv_imagep [simp]: "inv_imagep r f x y = r (f x) (f y)"
@@ -1115,8 +1021,7 @@
 subsubsection \<open>Powerset\<close>
 
 definition Powp :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
-where
-  "Powp A = (\<lambda>B. \<forall>x \<in> B. A x)"
+  where "Powp A = (\<lambda>B. \<forall>x \<in> B. A x)"
 
 lemma Powp_Pow_eq [pred_set_conv]: "Powp (\<lambda>x. x \<in> A) = (\<lambda>x. x \<in> Pow A)"
   by (auto simp add: Powp_def fun_eq_iff)
@@ -1130,28 +1035,31 @@
   assumes "finite A"
   shows "Id_on A = Finite_Set.fold (\<lambda>x. Set.insert (Pair x x)) {} A"
 proof -
-  interpret comp_fun_commute "\<lambda>x. Set.insert (Pair x x)" by standard auto
-  show ?thesis using assms unfolding Id_on_def by (induct A) simp_all
+  interpret comp_fun_commute "\<lambda>x. Set.insert (Pair x x)"
+    by standard auto
+  from assms show ?thesis
+    unfolding Id_on_def by (induct A) simp_all
 qed
 
 lemma comp_fun_commute_Image_fold:
   "comp_fun_commute (\<lambda>(x,y) A. if x \<in> S then Set.insert y A else A)"
 proof -
   interpret comp_fun_idem Set.insert
-      by (fact comp_fun_idem_insert)
-  show ?thesis 
-  by standard (auto simp add: fun_eq_iff comp_fun_commute split:prod.split)
+    by (fact comp_fun_idem_insert)
+  show ?thesis
+    by standard (auto simp add: fun_eq_iff comp_fun_commute split: prod.split)
 qed
 
 lemma Image_fold:
   assumes "finite R"
   shows "R `` S = Finite_Set.fold (\<lambda>(x,y) A. if x \<in> S then Set.insert y A else A) {} R"
 proof -
-  interpret comp_fun_commute "(\<lambda>(x,y) A. if x \<in> S then Set.insert y A else A)" 
+  interpret comp_fun_commute "(\<lambda>(x,y) A. if x \<in> S then Set.insert y A else A)"
     by (rule comp_fun_commute_Image_fold)
   have *: "\<And>x F. Set.insert x F `` S = (if fst x \<in> S then Set.insert (snd x) (F `` S) else (F `` S))"
     by (force intro: rev_ImageI)
-  show ?thesis using assms by (induct R) (auto simp: *)
+  show ?thesis
+    using assms by (induct R) (auto simp: *)
 qed
 
 lemma insert_relcomp_union_fold:
@@ -1159,56 +1067,56 @@
   shows "{x} O S \<union> X = Finite_Set.fold (\<lambda>(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') X S"
 proof -
   interpret comp_fun_commute "\<lambda>(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A'"
-  proof - 
-    interpret comp_fun_idem Set.insert by (fact comp_fun_idem_insert)
+  proof -
+    interpret comp_fun_idem Set.insert
+      by (fact comp_fun_idem_insert)
     show "comp_fun_commute (\<lambda>(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A')"
-    by standard (auto simp add: fun_eq_iff split:prod.split)
+      by standard (auto simp add: fun_eq_iff split: prod.split)
   qed
-  have *: "{x} O S = {(x', z). x' = fst x \<and> (snd x,z) \<in> S}" by (auto simp: relcomp_unfold intro!: exI)
-  show ?thesis unfolding *
-  using \<open>finite S\<close> by (induct S) (auto split: prod.split)
+  have *: "{x} O S = {(x', z). x' = fst x \<and> (snd x, z) \<in> S}"
+    by (auto simp: relcomp_unfold intro!: exI)
+  show ?thesis
+    unfolding * using \<open>finite S\<close> by (induct S) (auto split: prod.split)
 qed
 
 lemma insert_relcomp_fold:
   assumes "finite S"
-  shows "Set.insert x R O S = 
+  shows "Set.insert x R O S =
     Finite_Set.fold (\<lambda>(w,z) A'. if snd x = w then Set.insert (fst x,z) A' else A') (R O S) S"
 proof -
-  have "Set.insert x R O S = ({x} O S) \<union> (R O S)" by auto
-  then show ?thesis by (auto simp: insert_relcomp_union_fold[OF assms])
+  have "Set.insert x R O S = ({x} O S) \<union> (R O S)"
+    by auto
+  then show ?thesis
+    by (auto simp: insert_relcomp_union_fold [OF assms])
 qed
 
 lemma comp_fun_commute_relcomp_fold:
   assumes "finite S"
-  shows "comp_fun_commute (\<lambda>(x,y) A. 
+  shows "comp_fun_commute (\<lambda>(x,y) A.
     Finite_Set.fold (\<lambda>(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S)"
 proof -
-  have *: "\<And>a b A. 
+  have *: "\<And>a b A.
     Finite_Set.fold (\<lambda>(w, z) A'. if b = w then Set.insert (a, z) A' else A') A S = {(a,b)} O S \<union> A"
     by (auto simp: insert_relcomp_union_fold[OF assms] cong: if_cong)
-  show ?thesis by standard (auto simp: *)
+  show ?thesis
+    by standard (auto simp: *)
 qed
 
 lemma relcomp_fold:
-  assumes "finite R"
-  assumes "finite S"
-  shows "R O S = Finite_Set.fold 
+  assumes "finite R" "finite S"
+  shows "R O S = Finite_Set.fold
     (\<lambda>(x,y) A. Finite_Set.fold (\<lambda>(w,z) A'. if y = w then Set.insert (x,z) A' else A') A S) {} R"
-  using assms by (induct R)
+  using assms
+  by (induct R)
     (auto simp: comp_fun_commute.fold_insert comp_fun_commute_relcomp_fold insert_relcomp_fold
       cong: if_cong)
 
 text \<open>Misc\<close>
 
 abbreviation (input) transP :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> bool"
-where \<comment> \<open>FIXME drop\<close>
-  "transP r \<equiv> trans {(x, y). r x y}"
+  where "transP r \<equiv> trans {(x, y). r x y}"  (* FIXME drop *)
 
-abbreviation (input)
-  "RangeP \<equiv> Rangep"
-
-abbreviation (input)
-  "DomainP \<equiv> Domainp"
-
+abbreviation (input) "RangeP \<equiv> Rangep"
+abbreviation (input) "DomainP \<equiv> Domainp"
 
 end