src/ZF/func.thy
changeset 76213 e44d86131648
parent 71085 950e1cfe0fe4
child 76214 0c18df79b1c8
--- a/src/ZF/func.thy	Tue Sep 27 13:34:54 2022 +0200
+++ b/src/ZF/func.thy	Tue Sep 27 16:51:35 2022 +0100
@@ -9,11 +9,11 @@
 
 subsection\<open>The Pi Operator: Dependent Function Space\<close>
 
-lemma subset_Sigma_imp_relation: "r \<subseteq> Sigma(A,B) ==> relation(r)"
+lemma subset_Sigma_imp_relation: "r \<subseteq> Sigma(A,B) \<Longrightarrow> relation(r)"
 by (simp add: relation_def, blast)
 
 lemma relation_converse_converse [simp]:
-     "relation(r) ==> converse(converse(r)) = r"
+     "relation(r) \<Longrightarrow> converse(converse(r)) = r"
 by (simp add: relation_def, blast)
 
 lemma relation_restrict [simp]:  "relation(restrict(r,A))"
@@ -28,23 +28,23 @@
     "f \<in> Pi(A,B) \<longleftrightarrow> f<=Sigma(A,B) & (\<forall>x\<in>A. \<exists>!y. <x,y>: f)"
 by (unfold Pi_def function_def, blast)
 
-lemma fun_is_function: "f \<in> Pi(A,B) ==> function(f)"
+lemma fun_is_function: "f \<in> Pi(A,B) \<Longrightarrow> function(f)"
 by (simp only: Pi_iff)
 
 lemma function_imp_Pi:
-     "[|function(f); relation(f)|] ==> f \<in> domain(f) -> range(f)"
+     "\<lbrakk>function(f); relation(f)\<rbrakk> \<Longrightarrow> f \<in> domain(f) -> range(f)"
 by (simp add: Pi_iff relation_def, blast)
 
 lemma functionI:
-     "[| !!x y y'. [| <x,y>:r; <x,y'>:r |] ==> y=y' |] ==> function(r)"
+     "\<lbrakk>\<And>x y y'. \<lbrakk><x,y>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
 by (simp add: function_def, blast)
 
 (*Functions are relations*)
-lemma fun_is_rel: "f \<in> Pi(A,B) ==> f \<subseteq> Sigma(A,B)"
+lemma fun_is_rel: "f \<in> Pi(A,B) \<Longrightarrow> f \<subseteq> Sigma(A,B)"
 by (unfold Pi_def, blast)
 
 lemma Pi_cong:
-    "[| A=A';  !!x. x \<in> A' ==> B(x)=B'(x) |] ==> Pi(A,B) = Pi(A',B')"
+    "\<lbrakk>A=A';  \<And>x. x \<in> A' \<Longrightarrow> B(x)=B'(x)\<rbrakk> \<Longrightarrow> Pi(A,B) = Pi(A',B')"
 by (simp add: Pi_def cong add: Sigma_cong)
 
 (*Sigma_cong, Pi_cong NOT given to Addcongs: they cause
@@ -52,57 +52,57 @@
   Sigmas and Pis are abbreviated as * or -> *)
 
 (*Weakening one function type to another; see also Pi_type*)
-lemma fun_weaken_type: "[| f \<in> A->B;  B<=D |] ==> f \<in> A->D"
+lemma fun_weaken_type: "\<lbrakk>f \<in> A->B;  B<=D\<rbrakk> \<Longrightarrow> f \<in> A->D"
 by (unfold Pi_def, best)
 
 subsection\<open>Function Application\<close>
 
-lemma apply_equality2: "[| <a,b>: f;  <a,c>: f;  f \<in> Pi(A,B) |] ==> b=c"
+lemma apply_equality2: "\<lbrakk><a,b>: f;  <a,c>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b=c"
 by (unfold Pi_def function_def, blast)
 
-lemma function_apply_equality: "[| <a,b>: f;  function(f) |] ==> f`a = b"
+lemma function_apply_equality: "\<lbrakk><a,b>: f;  function(f)\<rbrakk> \<Longrightarrow> f`a = b"
 by (unfold apply_def function_def, blast)
 
-lemma apply_equality: "[| <a,b>: f;  f \<in> Pi(A,B) |] ==> f`a = b"
+lemma apply_equality: "\<lbrakk><a,b>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> f`a = b"
 apply (unfold Pi_def)
 apply (blast intro: function_apply_equality)
 done
 
 (*Applying a function outside its domain yields 0*)
-lemma apply_0: "a \<notin> domain(f) ==> f`a = 0"
+lemma apply_0: "a \<notin> domain(f) \<Longrightarrow> f`a = 0"
 by (unfold apply_def, blast)
 
-lemma Pi_memberD: "[| f \<in> Pi(A,B);  c \<in> f |] ==> \<exists>x\<in>A.  c = <x,f`x>"
+lemma Pi_memberD: "\<lbrakk>f \<in> Pi(A,B);  c \<in> f\<rbrakk> \<Longrightarrow> \<exists>x\<in>A.  c = <x,f`x>"
 apply (frule fun_is_rel)
 apply (blast dest: apply_equality)
 done
 
-lemma function_apply_Pair: "[| function(f);  a \<in> domain(f)|] ==> <a,f`a>: f"
+lemma function_apply_Pair: "\<lbrakk>function(f);  a \<in> domain(f)\<rbrakk> \<Longrightarrow> <a,f`a>: f"
 apply (simp add: function_def, clarify)
 apply (subgoal_tac "f`a = y", blast)
 apply (simp add: apply_def, blast)
 done
 
-lemma apply_Pair: "[| f \<in> Pi(A,B);  a \<in> A |] ==> <a,f`a>: f"
+lemma apply_Pair: "\<lbrakk>f \<in> Pi(A,B);  a \<in> A\<rbrakk> \<Longrightarrow> <a,f`a>: f"
 apply (simp add: Pi_iff)
 apply (blast intro: function_apply_Pair)
 done
 
 (*Conclusion is flexible -- use rule_tac or else apply_funtype below!*)
-lemma apply_type [TC]: "[| f \<in> Pi(A,B);  a \<in> A |] ==> f`a \<in> B(a)"
+lemma apply_type [TC]: "\<lbrakk>f \<in> Pi(A,B);  a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> B(a)"
 by (blast intro: apply_Pair dest: fun_is_rel)
 
 (*This version is acceptable to the simplifier*)
-lemma apply_funtype: "[| f \<in> A->B;  a \<in> A |] ==> f`a \<in> B"
+lemma apply_funtype: "\<lbrakk>f \<in> A->B;  a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> B"
 by (blast dest: apply_type)
 
-lemma apply_iff: "f \<in> Pi(A,B) ==> <a,b>: f \<longleftrightarrow> a \<in> A & f`a = b"
+lemma apply_iff: "f \<in> Pi(A,B) \<Longrightarrow> <a,b>: f \<longleftrightarrow> a \<in> A & f`a = b"
 apply (frule fun_is_rel)
 apply (blast intro!: apply_Pair apply_equality)
 done
 
 (*Refining one Pi type to another*)
-lemma Pi_type: "[| f \<in> Pi(A,C);  !!x. x \<in> A ==> f`x \<in> B(x) |] ==> f \<in> Pi(A,B)"
+lemma Pi_type: "\<lbrakk>f \<in> Pi(A,C);  \<And>x. x \<in> A \<Longrightarrow> f`x \<in> B(x)\<rbrakk> \<Longrightarrow> f \<in> Pi(A,B)"
 apply (simp only: Pi_iff)
 apply (blast dest: function_apply_equality)
 done
@@ -114,38 +114,38 @@
 by (blast intro: Pi_type dest: apply_type)
 
 lemma Pi_weaken_type:
-        "[| f \<in> Pi(A,B);  !!x. x \<in> A ==> B(x)<=C(x) |] ==> f \<in> Pi(A,C)"
+        "\<lbrakk>f \<in> Pi(A,B);  \<And>x. x \<in> A \<Longrightarrow> B(x)<=C(x)\<rbrakk> \<Longrightarrow> f \<in> Pi(A,C)"
 by (blast intro: Pi_type dest: apply_type)
 
 
 (** Elimination of membership in a function **)
 
-lemma domain_type: "[| <a,b> \<in> f;  f \<in> Pi(A,B) |] ==> a \<in> A"
+lemma domain_type: "\<lbrakk><a,b> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A"
 by (blast dest: fun_is_rel)
 
-lemma range_type: "[| <a,b> \<in> f;  f \<in> Pi(A,B) |] ==> b \<in> B(a)"
+lemma range_type: "\<lbrakk><a,b> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b \<in> B(a)"
 by (blast dest: fun_is_rel)
 
-lemma Pair_mem_PiD: "[| <a,b>: f;  f \<in> Pi(A,B) |] ==> a \<in> A & b \<in> B(a) & f`a = b"
+lemma Pair_mem_PiD: "\<lbrakk><a,b>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A & b \<in> B(a) & f`a = b"
 by (blast intro: domain_type range_type apply_equality)
 
 subsection\<open>Lambda Abstraction\<close>
 
-lemma lamI: "a \<in> A ==> <a,b(a)> \<in> (\<lambda>x\<in>A. b(x))"
+lemma lamI: "a \<in> A \<Longrightarrow> <a,b(a)> \<in> (\<lambda>x\<in>A. b(x))"
 apply (unfold lam_def)
 apply (erule RepFunI)
 done
 
 lemma lamE:
-    "[| p: (\<lambda>x\<in>A. b(x));  !!x.[| x \<in> A; p=<x,b(x)> |] ==> P
-     |] ==>  P"
+    "\<lbrakk>p: (\<lambda>x\<in>A. b(x));  \<And>x.\<lbrakk>x \<in> A; p=<x,b(x)>\<rbrakk> \<Longrightarrow> P
+\<rbrakk> \<Longrightarrow>  P"
 by (simp add: lam_def, blast)
 
-lemma lamD: "[| <a,c>: (\<lambda>x\<in>A. b(x)) |] ==> c = b(a)"
+lemma lamD: "\<lbrakk><a,c>: (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> c = b(a)"
 by (simp add: lam_def)
 
 lemma lam_type [TC]:
-    "[| !!x. x \<in> A ==> b(x): B(x) |] ==> (\<lambda>x\<in>A. b(x)) \<in> Pi(A,B)"
+    "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> b(x): B(x)\<rbrakk> \<Longrightarrow> (\<lambda>x\<in>A. b(x)) \<in> Pi(A,B)"
 by (simp add: lam_def Pi_def function_def, blast)
 
 lemma lam_funtype: "(\<lambda>x\<in>A. b(x)) \<in> A -> {b(x). x \<in> A}"
@@ -160,7 +160,7 @@
 lemma beta_if [simp]: "(\<lambda>x\<in>A. b(x)) ` a = (if a \<in> A then b(a) else 0)"
 by (simp add: apply_def lam_def, blast)
 
-lemma beta: "a \<in> A ==> (\<lambda>x\<in>A. b(x)) ` a = b(a)"
+lemma beta: "a \<in> A \<Longrightarrow> (\<lambda>x\<in>A. b(x)) ` a = b(a)"
 by (simp add: apply_def lam_def, blast)
 
 lemma lam_empty [simp]: "(\<lambda>x\<in>0. b(x)) = 0"
@@ -171,17 +171,17 @@
 
 (*congruence rule for lambda abstraction*)
 lemma lam_cong [cong]:
-    "[| A=A';  !!x. x \<in> A' ==> b(x)=b'(x) |] ==> Lambda(A,b) = Lambda(A',b')"
+    "\<lbrakk>A=A';  \<And>x. x \<in> A' \<Longrightarrow> b(x)=b'(x)\<rbrakk> \<Longrightarrow> Lambda(A,b) = Lambda(A',b')"
 by (simp only: lam_def cong add: RepFun_cong)
 
 lemma lam_theI:
-    "(!!x. x \<in> A ==> \<exists>!y. Q(x,y)) ==> \<exists>f. \<forall>x\<in>A. Q(x, f`x)"
+    "(\<And>x. x \<in> A \<Longrightarrow> \<exists>!y. Q(x,y)) \<Longrightarrow> \<exists>f. \<forall>x\<in>A. Q(x, f`x)"
 apply (rule_tac x = "\<lambda>x\<in>A. THE y. Q (x,y)" in exI)
 apply simp
 apply (blast intro: theI)
 done
 
-lemma lam_eqE: "[| (\<lambda>x\<in>A. f(x)) = (\<lambda>x\<in>A. g(x));  a \<in> A |] ==> f(a)=g(a)"
+lemma lam_eqE: "\<lbrakk>(\<lambda>x\<in>A. f(x)) = (\<lambda>x\<in>A. g(x));  a \<in> A\<rbrakk> \<Longrightarrow> f(a)=g(a)"
 by (fast intro!: lamI elim: equalityE lamE)
 
 
@@ -207,26 +207,26 @@
 (*Semi-extensionality!*)
 
 lemma fun_subset:
-    "[| f \<in> Pi(A,B);  g \<in> Pi(C,D);  A<=C;
-        !!x. x \<in> A ==> f`x = g`x       |] ==> f<=g"
+    "\<lbrakk>f \<in> Pi(A,B);  g \<in> Pi(C,D);  A<=C;
+        \<And>x. x \<in> A \<Longrightarrow> f`x = g`x\<rbrakk> \<Longrightarrow> f<=g"
 by (force dest: Pi_memberD intro: apply_Pair)
 
 lemma fun_extension:
-    "[| f \<in> Pi(A,B);  g \<in> Pi(A,D);
-        !!x. x \<in> A ==> f`x = g`x       |] ==> f=g"
+    "\<lbrakk>f \<in> Pi(A,B);  g \<in> Pi(A,D);
+        \<And>x. x \<in> A \<Longrightarrow> f`x = g`x\<rbrakk> \<Longrightarrow> f=g"
 by (blast del: subsetI intro: subset_refl sym fun_subset)
 
-lemma eta [simp]: "f \<in> Pi(A,B) ==> (\<lambda>x\<in>A. f`x) = f"
+lemma eta [simp]: "f \<in> Pi(A,B) \<Longrightarrow> (\<lambda>x\<in>A. f`x) = f"
 apply (rule fun_extension)
 apply (auto simp add: lam_type apply_type beta)
 done
 
 lemma fun_extension_iff:
-     "[| f \<in> Pi(A,B); g \<in> Pi(A,C) |] ==> (\<forall>a\<in>A. f`a = g`a) \<longleftrightarrow> f=g"
+     "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(A,C)\<rbrakk> \<Longrightarrow> (\<forall>a\<in>A. f`a = g`a) \<longleftrightarrow> f=g"
 by (blast intro: fun_extension)
 
 (*thm by Mark Staples, proof by lcp*)
-lemma fun_subset_eq: "[| f \<in> Pi(A,B); g \<in> Pi(A,C) |] ==> f \<subseteq> g \<longleftrightarrow> (f = g)"
+lemma fun_subset_eq: "\<lbrakk>f \<in> Pi(A,B); g \<in> Pi(A,C)\<rbrakk> \<Longrightarrow> f \<subseteq> g \<longleftrightarrow> (f = g)"
 by (blast dest: apply_Pair
           intro: fun_extension apply_equality [symmetric])
 
@@ -234,7 +234,7 @@
 (*Every element of Pi(A,B) may be expressed as a lambda abstraction!*)
 lemma Pi_lamE:
   assumes major: "f \<in> Pi(A,B)"
-      and minor: "!!b. [| \<forall>x\<in>A. b(x):B(x);  f = (\<lambda>x\<in>A. b(x)) |] ==> P"
+      and minor: "\<And>b. \<lbrakk>\<forall>x\<in>A. b(x):B(x);  f = (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> P"
   shows "P"
 apply (rule minor)
 apply (rule_tac [2] eta [symmetric])
@@ -244,12 +244,12 @@
 
 subsection\<open>Images of Functions\<close>
 
-lemma image_lam: "C \<subseteq> A ==> (\<lambda>x\<in>A. b(x)) `` C = {b(x). x \<in> C}"
+lemma image_lam: "C \<subseteq> A \<Longrightarrow> (\<lambda>x\<in>A. b(x)) `` C = {b(x). x \<in> C}"
 by (unfold lam_def, blast)
 
 lemma Repfun_function_if:
      "function(f)
-      ==> {f`x. x \<in> C} = (if C \<subseteq> domain(f) then f``C else cons(0,f``C))"
+      \<Longrightarrow> {f`x. x \<in> C} = (if C \<subseteq> domain(f) then f``C else cons(0,f``C))"
 apply simp
 apply (intro conjI impI)
  apply (blast dest: function_apply_equality intro: function_apply_Pair)
@@ -261,10 +261,10 @@
 (*For this lemma and the next, the right-hand side could equivalently
   be written \<Union>x\<in>C. {f`x} *)
 lemma image_function:
-     "[| function(f);  C \<subseteq> domain(f) |] ==> f``C = {f`x. x \<in> C}"
+     "\<lbrakk>function(f);  C \<subseteq> domain(f)\<rbrakk> \<Longrightarrow> f``C = {f`x. x \<in> C}"
 by (simp add: Repfun_function_if)
 
-lemma image_fun: "[| f \<in> Pi(A,B);  C \<subseteq> A |] ==> f``C = {f`x. x \<in> C}"
+lemma image_fun: "\<lbrakk>f \<in> Pi(A,B);  C \<subseteq> A\<rbrakk> \<Longrightarrow> f``C = {f`x. x \<in> C}"
 apply (simp add: Pi_iff)
 apply (blast intro: image_function)
 done
@@ -274,7 +274,7 @@
 by (auto simp add: image_fun [OF f])
 
 lemma Pi_image_cons:
-     "[| f \<in> Pi(A,B);  x \<in> A |] ==> f `` cons(x,y) = cons(f`x, f``y)"
+     "\<lbrakk>f \<in> Pi(A,B);  x \<in> A\<rbrakk> \<Longrightarrow> f `` cons(x,y) = cons(f`x, f``y)"
 by (blast dest: apply_equality apply_Pair)
 
 
@@ -284,10 +284,10 @@
 by (unfold restrict_def, blast)
 
 lemma function_restrictI:
-    "function(f) ==> function(restrict(f,A))"
+    "function(f) \<Longrightarrow> function(restrict(f,A))"
 by (unfold restrict_def function_def, blast)
 
-lemma restrict_type2: "[| f \<in> Pi(C,B);  A<=C |] ==> restrict(f,A) \<in> Pi(A,B)"
+lemma restrict_type2: "\<lbrakk>f \<in> Pi(C,B);  A<=C\<rbrakk> \<Longrightarrow> restrict(f,A) \<in> Pi(A,B)"
 by (simp add: Pi_iff function_def restrict_def, blast)
 
 lemma restrict: "restrict(f,A) ` a = (if a \<in> A then f`a else 0)"
@@ -308,13 +308,13 @@
 apply (auto simp add: domain_def)
 done
 
-lemma restrict_idem: "f \<subseteq> Sigma(A,B) ==> restrict(f,A) = f"
+lemma restrict_idem: "f \<subseteq> Sigma(A,B) \<Longrightarrow> restrict(f,A) = f"
 by (simp add: restrict_def, blast)
 
 
 (*converse probably holds too*)
 lemma domain_restrict_idem:
-     "[| domain(r) \<subseteq> A; relation(r) |] ==> restrict(r,A) = r"
+     "\<lbrakk>domain(r) \<subseteq> A; relation(r)\<rbrakk> \<Longrightarrow> restrict(r,A) = r"
 by (simp add: restrict_def relation_def, blast)
 
 lemma domain_restrict_lam [simp]: "domain(restrict(Lambda(A,f),C)) = A \<inter> C"
@@ -327,11 +327,11 @@
 by (simp add: restrict apply_0)
 
 lemma restrict_lam_eq:
-    "A<=C ==> restrict(\<lambda>x\<in>C. b(x), A) = (\<lambda>x\<in>A. b(x))"
+    "A<=C \<Longrightarrow> restrict(\<lambda>x\<in>C. b(x), A) = (\<lambda>x\<in>A. b(x))"
 by (unfold restrict_def lam_def, auto)
 
 lemma fun_cons_restrict_eq:
-     "f \<in> cons(a, b) -> B ==> f = cons(<a, f ` a>, restrict(f, b))"
+     "f \<in> cons(a, b) -> B \<Longrightarrow> f = cons(<a, f ` a>, restrict(f, b))"
 apply (rule equalityI)
  prefer 2 apply (blast intro: apply_Pair restrict_subset [THEN subsetD])
 apply (auto dest!: Pi_memberD simp add: restrict_def lam_def)
@@ -343,14 +343,14 @@
 (** The Union of a set of COMPATIBLE functions is a function **)
 
 lemma function_Union:
-    "[| \<forall>x\<in>S. function(x);
-        \<forall>x\<in>S. \<forall>y\<in>S. x<=y | y<=x  |]
-     ==> function(\<Union>(S))"
+    "\<lbrakk>\<forall>x\<in>S. function(x);
+        \<forall>x\<in>S. \<forall>y\<in>S. x<=y | y<=x\<rbrakk>
+     \<Longrightarrow> function(\<Union>(S))"
 by (unfold function_def, blast)
 
 lemma fun_Union:
-    "[| \<forall>f\<in>S. \<exists>C D. f \<in> C->D;
-             \<forall>f\<in>S. \<forall>y\<in>S. f<=y | y<=f  |] ==>
+    "\<lbrakk>\<forall>f\<in>S. \<exists>C D. f \<in> C->D;
+             \<forall>f\<in>S. \<forall>y\<in>S. f<=y | y<=f\<rbrakk> \<Longrightarrow>
           \<Union>(S) \<in> domain(\<Union>(S)) -> range(\<Union>(S))"
 apply (unfold Pi_def)
 apply (blast intro!: rel_Union function_Union)
@@ -368,48 +368,48 @@
                 subset_trans [OF _ Un_upper2]
 
 lemma fun_disjoint_Un:
-     "[| f \<in> A->B;  g \<in> C->D;  A \<inter> C = 0  |]
-      ==> (f \<union> g) \<in> (A \<union> C) -> (B \<union> D)"
+     "\<lbrakk>f \<in> A->B;  g \<in> C->D;  A \<inter> C = 0\<rbrakk>
+      \<Longrightarrow> (f \<union> g) \<in> (A \<union> C) -> (B \<union> D)"
 (*Prove the product and domain subgoals using distributive laws*)
 apply (simp add: Pi_iff extension Un_rls)
 apply (unfold function_def, blast)
 done
 
-lemma fun_disjoint_apply1: "a \<notin> domain(g) ==> (f \<union> g)`a = f`a"
+lemma fun_disjoint_apply1: "a \<notin> domain(g) \<Longrightarrow> (f \<union> g)`a = f`a"
 by (simp add: apply_def, blast)
 
-lemma fun_disjoint_apply2: "c \<notin> domain(f) ==> (f \<union> g)`c = g`c"
+lemma fun_disjoint_apply2: "c \<notin> domain(f) \<Longrightarrow> (f \<union> g)`c = g`c"
 by (simp add: apply_def, blast)
 
 subsection\<open>Domain and Range of a Function or Relation\<close>
 
-lemma domain_of_fun: "f \<in> Pi(A,B) ==> domain(f)=A"
+lemma domain_of_fun: "f \<in> Pi(A,B) \<Longrightarrow> domain(f)=A"
 by (unfold Pi_def, blast)
 
-lemma apply_rangeI: "[| f \<in> Pi(A,B);  a \<in> A |] ==> f`a \<in> range(f)"
+lemma apply_rangeI: "\<lbrakk>f \<in> Pi(A,B);  a \<in> A\<rbrakk> \<Longrightarrow> f`a \<in> range(f)"
 by (erule apply_Pair [THEN rangeI], assumption)
 
-lemma range_of_fun: "f \<in> Pi(A,B) ==> f \<in> A->range(f)"
+lemma range_of_fun: "f \<in> Pi(A,B) \<Longrightarrow> f \<in> A->range(f)"
 by (blast intro: Pi_type apply_rangeI)
 
 subsection\<open>Extensions of Functions\<close>
 
 lemma fun_extend:
-     "[| f \<in> A->B;  c\<notin>A |] ==> cons(<c,b>,f) \<in> cons(c,A) -> cons(b,B)"
+     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(<c,b>,f) \<in> cons(c,A) -> cons(b,B)"
 apply (frule singleton_fun [THEN fun_disjoint_Un], blast)
 apply (simp add: cons_eq)
 done
 
 lemma fun_extend3:
-     "[| f \<in> A->B;  c\<notin>A;  b \<in> B |] ==> cons(<c,b>,f) \<in> cons(c,A) -> B"
+     "\<lbrakk>f \<in> A->B;  c\<notin>A;  b \<in> B\<rbrakk> \<Longrightarrow> cons(<c,b>,f) \<in> cons(c,A) -> B"
 by (blast intro: fun_extend [THEN fun_weaken_type])
 
 lemma extend_apply:
-     "c \<notin> domain(f) ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+     "c \<notin> domain(f) \<Longrightarrow> cons(<c,b>,f)`a = (if a=c then b else f`a)"
 by (auto simp add: apply_def)
 
 lemma fun_extend_apply [simp]:
-     "[| f \<in> A->B;  c\<notin>A |] ==> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(<c,b>,f)`a = (if a=c then b else f`a)"
 apply (rule extend_apply)
 apply (simp add: Pi_def, blast)
 done
@@ -418,7 +418,7 @@
 
 (*For Finite.ML.  Inclusion of right into left is easy*)
 lemma cons_fun_eq:
-     "c \<notin> A ==> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(<c,b>, f)})"
+     "c \<notin> A \<Longrightarrow> cons(c,A) -> B = (\<Union>f \<in> A->B. \<Union>b\<in>B. {cons(<c,b>, f)})"
 apply (rule equalityI)
 apply (safe elim!: fun_extend3)
 (*Inclusion of left into right*)
@@ -443,7 +443,7 @@
 
 definition
   update  :: "[i,i,i] => i"  where
-   "update(f,a,b) == \<lambda>x\<in>cons(a, domain(f)). if(x=a, b, f`x)"
+   "update(f,a,b) \<equiv> \<lambda>x\<in>cons(a, domain(f)). if(x=a, b, f`x)"
 
 nonterminal updbinds and updbind
 
@@ -467,7 +467,7 @@
 apply (simp_all add: apply_0)
 done
 
-lemma update_idem: "[| f`x = y;  f \<in> Pi(A,B);  x \<in> A |] ==> f(x:=y) = f"
+lemma update_idem: "\<lbrakk>f`x = y;  f \<in> Pi(A,B);  x \<in> A\<rbrakk> \<Longrightarrow> f(x:=y) = f"
 apply (unfold update_def)
 apply (simp add: domain_of_fun cons_absorb)
 apply (rule fun_extension)
@@ -475,13 +475,13 @@
 done
 
 
-(* [| f \<in> Pi(A, B); x \<in> A |] ==> f(x := f`x) = f *)
+(* \<lbrakk>f \<in> Pi(A, B); x \<in> A\<rbrakk> \<Longrightarrow> f(x := f`x) = f *)
 declare refl [THEN update_idem, simp]
 
 lemma domain_update [simp]: "domain(f(x:=y)) = cons(x, domain(f))"
 by (unfold update_def, simp)
 
-lemma update_type: "[| f \<in> Pi(A,B);  x \<in> A;  y \<in> B(x) |] ==> f(x:=y) \<in> Pi(A, B)"
+lemma update_type: "\<lbrakk>f \<in> Pi(A,B);  x \<in> A;  y \<in> B(x)\<rbrakk> \<Longrightarrow> f(x:=y) \<in> Pi(A, B)"
 apply (unfold update_def)
 apply (simp add: domain_of_fun cons_absorb apply_funtype lam_type)
 done
@@ -493,96 +493,96 @@
 
 (*Not easy to express monotonicity in P, since any "bigger" predicate
   would have to be single-valued*)
-lemma Replace_mono: "A<=B ==> Replace(A,P) \<subseteq> Replace(B,P)"
+lemma Replace_mono: "A<=B \<Longrightarrow> Replace(A,P) \<subseteq> Replace(B,P)"
 by (blast elim!: ReplaceE)
 
-lemma RepFun_mono: "A<=B ==> {f(x). x \<in> A} \<subseteq> {f(x). x \<in> B}"
+lemma RepFun_mono: "A<=B \<Longrightarrow> {f(x). x \<in> A} \<subseteq> {f(x). x \<in> B}"
 by blast
 
-lemma Pow_mono: "A<=B ==> Pow(A) \<subseteq> Pow(B)"
+lemma Pow_mono: "A<=B \<Longrightarrow> Pow(A) \<subseteq> Pow(B)"
 by blast
 
-lemma Union_mono: "A<=B ==> \<Union>(A) \<subseteq> \<Union>(B)"
+lemma Union_mono: "A<=B \<Longrightarrow> \<Union>(A) \<subseteq> \<Union>(B)"
 by blast
 
 lemma UN_mono:
-    "[| A<=C;  !!x. x \<in> A ==> B(x)<=D(x) |] ==> (\<Union>x\<in>A. B(x)) \<subseteq> (\<Union>x\<in>C. D(x))"
+    "\<lbrakk>A<=C;  \<And>x. x \<in> A \<Longrightarrow> B(x)<=D(x)\<rbrakk> \<Longrightarrow> (\<Union>x\<in>A. B(x)) \<subseteq> (\<Union>x\<in>C. D(x))"
 by blast
 
 (*Intersection is ANTI-monotonic.  There are TWO premises! *)
-lemma Inter_anti_mono: "[| A<=B;  A\<noteq>0 |] ==> \<Inter>(B) \<subseteq> \<Inter>(A)"
+lemma Inter_anti_mono: "\<lbrakk>A<=B;  A\<noteq>0\<rbrakk> \<Longrightarrow> \<Inter>(B) \<subseteq> \<Inter>(A)"
 by blast
 
-lemma cons_mono: "C<=D ==> cons(a,C) \<subseteq> cons(a,D)"
+lemma cons_mono: "C<=D \<Longrightarrow> cons(a,C) \<subseteq> cons(a,D)"
 by blast
 
-lemma Un_mono: "[| A<=C;  B<=D |] ==> A \<union> B \<subseteq> C \<union> D"
+lemma Un_mono: "\<lbrakk>A<=C;  B<=D\<rbrakk> \<Longrightarrow> A \<union> B \<subseteq> C \<union> D"
 by blast
 
-lemma Int_mono: "[| A<=C;  B<=D |] ==> A \<inter> B \<subseteq> C \<inter> D"
+lemma Int_mono: "\<lbrakk>A<=C;  B<=D\<rbrakk> \<Longrightarrow> A \<inter> B \<subseteq> C \<inter> D"
 by blast
 
-lemma Diff_mono: "[| A<=C;  D<=B |] ==> A-B \<subseteq> C-D"
+lemma Diff_mono: "\<lbrakk>A<=C;  D<=B\<rbrakk> \<Longrightarrow> A-B \<subseteq> C-D"
 by blast
 
 subsubsection\<open>Standard Products, Sums and Function Spaces\<close>
 
 lemma Sigma_mono [rule_format]:
-     "[| A<=C;  !!x. x \<in> A \<longrightarrow> B(x) \<subseteq> D(x) |] ==> Sigma(A,B) \<subseteq> Sigma(C,D)"
+     "\<lbrakk>A<=C;  \<And>x. x \<in> A \<longrightarrow> B(x) \<subseteq> D(x)\<rbrakk> \<Longrightarrow> Sigma(A,B) \<subseteq> Sigma(C,D)"
 by blast
 
-lemma sum_mono: "[| A<=C;  B<=D |] ==> A+B \<subseteq> C+D"
+lemma sum_mono: "\<lbrakk>A<=C;  B<=D\<rbrakk> \<Longrightarrow> A+B \<subseteq> C+D"
 by (unfold sum_def, blast)
 
 (*Note that B->A and C->A are typically disjoint!*)
-lemma Pi_mono: "B<=C ==> A->B \<subseteq> A->C"
+lemma Pi_mono: "B<=C \<Longrightarrow> A->B \<subseteq> A->C"
 by (blast intro: lam_type elim: Pi_lamE)
 
-lemma lam_mono: "A<=B ==> Lambda(A,c) \<subseteq> Lambda(B,c)"
+lemma lam_mono: "A<=B \<Longrightarrow> Lambda(A,c) \<subseteq> Lambda(B,c)"
 apply (unfold lam_def)
 apply (erule RepFun_mono)
 done
 
 subsubsection\<open>Converse, Domain, Range, Field\<close>
 
-lemma converse_mono: "r<=s ==> converse(r) \<subseteq> converse(s)"
+lemma converse_mono: "r<=s \<Longrightarrow> converse(r) \<subseteq> converse(s)"
 by blast
 
-lemma domain_mono: "r<=s ==> domain(r)<=domain(s)"
+lemma domain_mono: "r<=s \<Longrightarrow> domain(r)<=domain(s)"
 by blast
 
 lemmas domain_rel_subset = subset_trans [OF domain_mono domain_subset]
 
-lemma range_mono: "r<=s ==> range(r)<=range(s)"
+lemma range_mono: "r<=s \<Longrightarrow> range(r)<=range(s)"
 by blast
 
 lemmas range_rel_subset = subset_trans [OF range_mono range_subset]
 
-lemma field_mono: "r<=s ==> field(r)<=field(s)"
+lemma field_mono: "r<=s \<Longrightarrow> field(r)<=field(s)"
 by blast
 
-lemma field_rel_subset: "r \<subseteq> A*A ==> field(r) \<subseteq> A"
+lemma field_rel_subset: "r \<subseteq> A*A \<Longrightarrow> field(r) \<subseteq> A"
 by (erule field_mono [THEN subset_trans], blast)
 
 
 subsubsection\<open>Images\<close>
 
 lemma image_pair_mono:
-    "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r``A \<subseteq> s``B"
+    "\<lbrakk>\<And>x y. <x,y>:r \<Longrightarrow> <x,y>:s;  A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
 by blast
 
 lemma vimage_pair_mono:
-    "[| !! x y. <x,y>:r ==> <x,y>:s;  A<=B |] ==> r-``A \<subseteq> s-``B"
+    "\<lbrakk>\<And>x y. <x,y>:r \<Longrightarrow> <x,y>:s;  A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
 by blast
 
-lemma image_mono: "[| r<=s;  A<=B |] ==> r``A \<subseteq> s``B"
+lemma image_mono: "\<lbrakk>r<=s;  A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
 by blast
 
-lemma vimage_mono: "[| r<=s;  A<=B |] ==> r-``A \<subseteq> s-``B"
+lemma vimage_mono: "\<lbrakk>r<=s;  A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
 by blast
 
 lemma Collect_mono:
-    "[| A<=B;  !!x. x \<in> A ==> P(x) \<longrightarrow> Q(x) |] ==> Collect(A,P) \<subseteq> Collect(B,Q)"
+    "\<lbrakk>A<=B;  \<And>x. x \<in> A \<Longrightarrow> P(x) \<longrightarrow> Q(x)\<rbrakk> \<Longrightarrow> Collect(A,P) \<subseteq> Collect(B,Q)"
 by blast
 
 (*Used in intr_elim.ML and in individual datatype definitions*)
@@ -592,7 +592,7 @@
 (* Useful with simp; contributed by Clemens Ballarin. *)
 
 lemma bex_image_simp:
-  "[| f \<in> Pi(X, Y); A \<subseteq> X |]  ==> (\<exists>x\<in>f``A. P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(f`x))"
+  "\<lbrakk>f \<in> Pi(X, Y); A \<subseteq> X\<rbrakk>  \<Longrightarrow> (\<exists>x\<in>f``A. P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(f`x))"
   apply safe
    apply rule
     prefer 2 apply assumption
@@ -601,7 +601,7 @@
   done
 
 lemma ball_image_simp:
-  "[| f \<in> Pi(X, Y); A \<subseteq> X |]  ==> (\<forall>x\<in>f``A. P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(f`x))"
+  "\<lbrakk>f \<in> Pi(X, Y); A \<subseteq> X\<rbrakk>  \<Longrightarrow> (\<forall>x\<in>f``A. P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(f`x))"
   apply safe
    apply (blast intro: apply_Pair)
   apply (drule bspec) apply assumption