More syntactic cleanup. LaTeX markup working
authorpaulson <lp15@cam.ac.uk>
Tue, 27 Sep 2022 17:46:52 +0100
changeset 76215 a642599ffdea
parent 76214 0c18df79b1c8
child 76216 9fc34f76b4e8
More syntactic cleanup. LaTeX markup working
src/ZF/AC.thy
src/ZF/AC/AC15_WO6.thy
src/ZF/AC/AC16_WO4.thy
src/ZF/AC/AC16_lemmas.thy
src/ZF/AC/AC18_AC19.thy
src/ZF/AC/AC7_AC9.thy
src/ZF/AC/AC_Equiv.thy
src/ZF/AC/Cardinal_aux.thy
src/ZF/AC/DC.thy
src/ZF/AC/HH.thy
src/ZF/AC/Hartog.thy
src/ZF/AC/WO2_AC16.thy
src/ZF/AC/WO6_WO1.thy
src/ZF/Arith.thy
src/ZF/Bin.thy
src/ZF/Bool.thy
src/ZF/Cardinal.thy
src/ZF/CardinalArith.thy
src/ZF/Cardinal_AC.thy
src/ZF/Coind/ECR.thy
src/ZF/Coind/Language.thy
src/ZF/Coind/Map.thy
src/ZF/Coind/Static.thy
src/ZF/Coind/Types.thy
src/ZF/Coind/Values.thy
src/ZF/Constructible/AC_in_L.thy
src/ZF/Constructible/DPow_absolute.thy
src/ZF/Constructible/Datatype_absolute.thy
src/ZF/Constructible/Formula.thy
src/ZF/Constructible/Internalize.thy
src/ZF/Constructible/L_axioms.thy
src/ZF/Constructible/Normal.thy
src/ZF/Constructible/Rank.thy
src/ZF/Constructible/Rank_Separation.thy
src/ZF/Constructible/Rec_Separation.thy
src/ZF/Constructible/Reflection.thy
src/ZF/Constructible/Relative.thy
src/ZF/Constructible/Satisfies_absolute.thy
src/ZF/Constructible/WF_absolute.thy
src/ZF/Constructible/WFrec.thy
src/ZF/Constructible/Wellorderings.thy
src/ZF/Epsilon.thy
src/ZF/EquivClass.thy
src/ZF/Finite.thy
src/ZF/Fixedpt.thy
src/ZF/IMP/Com.thy
src/ZF/IMP/Denotation.thy
src/ZF/IMP/Equiv.thy
src/ZF/Induct/Acc.thy
src/ZF/Induct/Binary_Trees.thy
src/ZF/Induct/Brouwer.thy
src/ZF/Induct/Comb.thy
src/ZF/Induct/Datatypes.thy
src/ZF/Induct/FoldSet.thy
src/ZF/Induct/ListN.thy
src/ZF/Induct/Multiset.thy
src/ZF/Induct/Mutil.thy
src/ZF/Induct/Ntree.thy
src/ZF/Induct/Primrec.thy
src/ZF/Induct/PropLog.thy
src/ZF/Induct/Rmap.thy
src/ZF/Induct/Term.thy
src/ZF/Induct/Tree_Forest.thy
src/ZF/Int.thy
src/ZF/IntDiv.thy
src/ZF/List.thy
src/ZF/Nat.thy
src/ZF/OrdQuant.thy
src/ZF/Order.thy
src/ZF/OrderArith.thy
src/ZF/OrderType.thy
src/ZF/Ordinal.thy
src/ZF/Perm.thy
src/ZF/QPair.thy
src/ZF/QUniv.thy
src/ZF/Resid/Confluence.thy
src/ZF/Resid/Redex.thy
src/ZF/Resid/Reduction.thy
src/ZF/Resid/Residuals.thy
src/ZF/Resid/Substitution.thy
src/ZF/Sum.thy
src/ZF/Trancl.thy
src/ZF/UNITY/AllocBase.thy
src/ZF/UNITY/AllocImpl.thy
src/ZF/UNITY/ClientImpl.thy
src/ZF/UNITY/Comp.thy
src/ZF/UNITY/Constrains.thy
src/ZF/UNITY/Distributor.thy
src/ZF/UNITY/FP.thy
src/ZF/UNITY/Follows.thy
src/ZF/UNITY/GenPrefix.thy
src/ZF/UNITY/Guar.thy
src/ZF/UNITY/Increasing.thy
src/ZF/UNITY/Merge.thy
src/ZF/UNITY/Monotonicity.thy
src/ZF/UNITY/MultisetSum.thy
src/ZF/UNITY/Mutex.thy
src/ZF/UNITY/State.thy
src/ZF/UNITY/SubstAx.thy
src/ZF/UNITY/UNITY.thy
src/ZF/UNITY/Union.thy
src/ZF/UNITY/WFair.thy
src/ZF/Univ.thy
src/ZF/WF.thy
src/ZF/ZF.thy
src/ZF/ZF_Base.thy
src/ZF/Zorn.thy
src/ZF/equalities.thy
src/ZF/ex/CoUnit.thy
src/ZF/ex/Commutation.thy
src/ZF/ex/Group.thy
src/ZF/ex/LList.thy
src/ZF/ex/Limit.thy
src/ZF/ex/NatSum.thy
src/ZF/ex/Primes.thy
src/ZF/ex/Ramsey.thy
src/ZF/ex/Ring.thy
src/ZF/ex/misc.thy
src/ZF/func.thy
src/ZF/pair.thy
src/ZF/upair.thy
--- a/src/ZF/AC.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -26,13 +26,13 @@
 done
 
 lemma AC_Pi_Pow: "\<exists>f. f \<in> (\<Prod>X \<in> Pow(C)-{0}. X)"
-apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE])
+apply (rule_tac B1 = "\<lambda>x. x" in AC_Pi [THEN exE])
 apply (erule_tac [2] exI, blast)
 done
 
 lemma AC_func:
      "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> (\<exists>y. y \<in> x)\<rbrakk> \<Longrightarrow> \<exists>f \<in> A->\<Union>(A). \<forall>x \<in> A. f`x \<in> x"
-apply (rule_tac B1 = "%x. x" in AC_Pi [THEN exE])
+apply (rule_tac B1 = "\<lambda>x. x" in AC_Pi [THEN exE])
 prefer 2 apply (blast dest: apply_type intro: Pi_type, blast)
 done
 
--- a/src/ZF/AC/AC15_WO6.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/AC15_WO6.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -34,7 +34,7 @@
 lemma lepoll_Sigma: "A\<noteq>0 \<Longrightarrow> B \<lesssim> A*B"
 apply (unfold lepoll_def)
 apply (erule not_emptyE)
-apply (rule_tac x = "\<lambda>z \<in> B. <x,z>" in exI)
+apply (rule_tac x = "\<lambda>z \<in> B. \<langle>x,z\<rangle>" in exI)
 apply (fast intro!: snd_conv lam_injective)
 done
 
@@ -69,7 +69,7 @@
 apply (erule exE)
 apply (rule_tac x = "\<lambda>x \<in> RepFun(A,P). \<mu> j. \<exists>a\<in>A. x=P(a) \<and> f`a=j" 
        in exI)
-apply (rule_tac d = "%y. P (converse (f) `y) " in lam_injective)
+apply (rule_tac d = "\<lambda>y. P (converse (f) `y) " in lam_injective)
 apply (erule RepFunE)
 apply (frule inj_is_fun [THEN apply_type], assumption)
 apply (fast intro: LeastI2 elim!: Ord_in_Ord inj_is_fun [THEN apply_type])
--- a/src/ZF/AC/AC16_WO4.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/AC16_WO4.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -79,7 +79,7 @@
 lemma succ_not_lepoll_lemma:
      "\<lbrakk>\<not>(\<exists>x \<in> A. f`x=y); f \<in> inj(A, B); y \<in> B\<rbrakk>   
       \<Longrightarrow> (\<lambda>a \<in> succ(A). if(a=A, y, f`a)) \<in> inj(succ(A), B)"
-apply (rule_tac d = "%z. if (z=y, A, converse (f) `z) " in lam_injective)
+apply (rule_tac d = "\<lambda>z. if (z=y, A, converse (f) `z) " in lam_injective)
 apply (force simp add: inj_is_fun [THEN apply_type])
 (*this preliminary simplification prevents looping somehow*)
 apply (simp (no_asm_simp))
--- a/src/ZF/AC/AC16_lemmas.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/AC16_lemmas.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -74,7 +74,7 @@
 apply (unfold lepoll_def)
 apply (rule_tac x = "\<lambda>y \<in> {y \<in> Pow(x) . y\<approx>succ (succ (n))}. 
                       <\<mu> i. i \<in> y, y-{\<mu> i. i \<in> y}>" in exI)
-apply (rule_tac d = "%z. cons (fst(z), snd(z))" in lam_injective)
+apply (rule_tac d = "\<lambda>z. cons (fst(z), snd(z))" in lam_injective)
  apply (blast intro!: Diff_sing_eqpoll intro: InfCard_Least_in)
 apply (simp, blast intro: InfCard_Least_in)
 done
@@ -147,7 +147,7 @@
 apply (unfold lepoll_def)
 apply (rule_tac x = "\<lambda>z \<in> {y\<in>Pow(x). y\<approx>succ(n)}. cons(succ(\<Union>(z)), z)" 
        in exI)
-apply (rule_tac d = "%z. z-{\<Union>(z) }" in lam_injective)
+apply (rule_tac d = "\<lambda>z. z-{\<Union>(z) }" in lam_injective)
 apply (blast intro!: succ_Union_in_x succ_Union_not_mem
              intro: cons_eqpoll_succ Ord_in_Ord
              dest!: InfCard_is_Card [THEN Card_is_Ord])
@@ -185,7 +185,7 @@
 apply (unfold eqpoll_def)
 apply (erule exE)
 apply (rule_tac x = "\<lambda>X \<in> {Y \<in> Pow (A) . \<exists>f. f \<in> bij (Y, n) }. f``X" in exI)
-apply (rule_tac d = "%Z. converse (f) ``Z" in lam_bijective)
+apply (rule_tac d = "\<lambda>Z. converse (f) ``Z" in lam_bijective)
 apply (fast intro!: bij_is_inj [THEN restrict_bij, THEN bij_converse_bij, 
                                 THEN comp_bij] 
             elim!: bij_is_fun [THEN fun_is_rel, THEN image_subset])
--- a/src/ZF/AC/AC18_AC19.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/AC18_AC19.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,7 +9,7 @@
 begin
 
 definition
-  uu    :: "i => i" where
+  uu    :: "i \<Rightarrow> i" where
     "uu(a) \<equiv> {c \<union> {0}. c \<in> a}"
 
 
@@ -48,7 +48,7 @@
 theorem (in AC18) AC19
 apply (unfold AC19_def)
 apply (intro allI impI)
-apply (rule AC18 [of _ "%x. x", THEN mp], blast)
+apply (rule AC18 [of _ "\<lambda>x. x", THEN mp], blast)
 done
 
 (* ********************************************************************** *)
--- a/src/ZF/AC/AC7_AC9.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/AC7_AC9.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -86,7 +86,7 @@
 (* ********************************************************************** *)
 
 lemma AC1_AC8_lemma1: 
-        "\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> \<and> B1 \<approx> B2   
+        "\<forall>B \<in> A. \<exists>B1 B2. B=\<langle>B1,B2\<rangle> \<and> B1 \<approx> B2   
         \<Longrightarrow> 0 \<notin> { bij(fst(B),snd(B)). B \<in> A }"
 apply (unfold eqpoll_def, auto)
 done
@@ -110,7 +110,7 @@
 
 lemma AC8_AC9_lemma:
      "\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1 \<approx> B2 
-      \<Longrightarrow> \<forall>B \<in> A*A. \<exists>B1 B2. B=<B1,B2> \<and> B1 \<approx> B2"
+      \<Longrightarrow> \<forall>B \<in> A*A. \<exists>B1 B2. B=\<langle>B1,B2\<rangle> \<and> B1 \<approx> B2"
 by fast
 
 lemma AC8_AC9: "AC8 \<Longrightarrow> AC9"
@@ -151,7 +151,7 @@
 lemma AC9_AC1_lemma2:
      "\<forall>B1 \<in> {(F*B)*N. B \<in> A} \<union> {cons(0,(F*B)*N). B \<in> A}.   
       \<forall>B2 \<in> {(F*B)*N. B \<in> A} \<union> {cons(0,(F*B)*N). B \<in> A}.   
-        f`<B1,B2> \<in> bij(B1, B2)   
+        f`\<langle>B1,B2\<rangle> \<in> bij(B1, B2)   
       \<Longrightarrow> (\<lambda>B \<in> A. snd(fst((f`<cons(0,(F*B)*N),(F*B)*N>)`0))) \<in> (\<Prod>X \<in> A. X)"
 apply (intro lam_type snd_type fst_type)
 apply (rule apply_type [OF _ consI1]) 
--- a/src/ZF/AC/AC_Equiv.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/AC_Equiv.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -46,11 +46,11 @@
 
 definition
 (* Auxiliary concepts needed below *)
-  pairwise_disjoint :: "i => o"  where
+  pairwise_disjoint :: "i \<Rightarrow> o"  where
     "pairwise_disjoint(A) \<equiv> \<forall>A1 \<in> A. \<forall>A2 \<in> A. A1 \<inter> A2 \<noteq> 0 \<longrightarrow> A1=A2"
 
 definition
-  sets_of_size_between :: "[i, i, i] => o"  where
+  sets_of_size_between :: "[i, i, i] \<Rightarrow> o"  where
     "sets_of_size_between(A,m,n) \<equiv> \<forall>B \<in> A. m \<lesssim> B \<and> B \<lesssim> n"
 
 
@@ -80,12 +80,12 @@
     "AC7 \<equiv> \<forall>A. 0\<notin>A \<and> (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow> (\<Prod>B \<in> A. B) \<noteq> 0"
 
 definition
-    "AC8 \<equiv> \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=<B1,B2> \<and> B1\<approx>B2)   
+    "AC8 \<equiv> \<forall>A. (\<forall>B \<in> A. \<exists>B1 B2. B=\<langle>B1,B2\<rangle> \<and> B1\<approx>B2)   
                    \<longrightarrow> (\<exists>f. \<forall>B \<in> A. f`B \<in> bij(fst(B),snd(B)))"
 
 definition
     "AC9 \<equiv> \<forall>A. (\<forall>B1 \<in> A. \<forall>B2 \<in> A. B1\<approx>B2) \<longrightarrow>   
-                   (\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`<B1,B2> \<in> bij(B1,B2))"
+                   (\<exists>f. \<forall>B1 \<in> A. \<forall>B2 \<in> A. f`\<langle>B1,B2\<rangle> \<in> bij(B1,B2))"
 
 definition
     "AC10(n) \<equiv>  \<forall>A. (\<forall>B \<in> A. \<not>Finite(B)) \<longrightarrow>   
@@ -140,21 +140,21 @@
 lemma rvimage_id: "rvimage(A,id(A),r) = r \<inter> A*A"
 apply (unfold rvimage_def)
 apply (rule equalityI, safe)
-apply (drule_tac P = "%a. <id (A) `xb,a>:r" in id_conv [THEN subst],
+apply (drule_tac P = "\<lambda>a. <id (A) `xb,a>:r" in id_conv [THEN subst],
        assumption)
-apply (drule_tac P = "%a. <a,ya>:r" in id_conv [THEN subst], (assumption+))
+apply (drule_tac P = "\<lambda>a. \<langle>a,ya\<rangle>:r" in id_conv [THEN subst], (assumption+))
 apply (fast intro: id_conv [THEN ssubst])
 done
 
 (* used only in Hartog.ML *)
 lemma ordertype_Int:
      "well_ord(A,r) \<Longrightarrow> ordertype(A, r \<inter> A*A) = ordertype(A,r)"
-apply (rule_tac P = "%a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst])
+apply (rule_tac P = "\<lambda>a. ordertype (A,a) =ordertype (A,r) " in rvimage_id [THEN subst])
 apply (erule id_bij [THEN bij_ordertype_vimage])
 done
 
 lemma lam_sing_bij: "(\<lambda>x \<in> A. {x}) \<in> bij(A, {{x}. x \<in> A})"
-apply (rule_tac d = "%z. THE x. z={x}" in lam_bijective)
+apply (rule_tac d = "\<lambda>z. THE x. z={x}" in lam_bijective)
 apply (auto simp add: the_equality)
 done
 
@@ -196,9 +196,9 @@
      "\<lbrakk>m \<in> nat; u \<lesssim> m\<rbrakk> \<Longrightarrow> domain(u) \<lesssim> m"
 apply (unfold lepoll_def)
 apply (erule exE)
-apply (rule_tac x = "\<lambda>x \<in> domain(u). \<mu> i. \<exists>y. <x,y> \<in> u \<and> f`<x,y> = i" 
+apply (rule_tac x = "\<lambda>x \<in> domain(u). \<mu> i. \<exists>y. \<langle>x,y\<rangle> \<in> u \<and> f`\<langle>x,y\<rangle> = i" 
        in exI)
-apply (rule_tac d = "%y. fst (converse(f) ` y) " in lam_injective)
+apply (rule_tac d = "\<lambda>y. fst (converse(f) ` y) " in lam_injective)
 apply (fast intro: LeastI2 nat_into_Ord [THEN Ord_in_Ord] 
                            inj_is_fun [THEN apply_type])
 apply (erule domainE)
--- a/src/ZF/AC/Cardinal_aux.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/Cardinal_aux.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -101,7 +101,7 @@
 lemma UN_sing_lepoll: "Ord(a) \<Longrightarrow> (\<Union>x \<in> a. {P(x)}) \<lesssim> a"
 apply (unfold lepoll_def)
 apply (rule_tac x = "\<lambda>z \<in> (\<Union>x \<in> a. {P (x) }) . (\<mu> i. P (i) =z) " in exI)
-apply (rule_tac d = "%z. P (z) " in lam_injective)
+apply (rule_tac d = "\<lambda>z. P (z) " in lam_injective)
 apply (fast intro!: Least_in_Ord)
 apply (fast intro: LeastI elim!: Ord_in_Ord)
 done
@@ -146,9 +146,9 @@
 apply (rule subsetI)
 apply (erule UN_E)
 apply (rule UN_I)
- apply (rule_tac P = "%z. x \<in> F (z) " in Least_in_Ord, (assumption+))
+ apply (rule_tac P = "\<lambda>z. x \<in> F (z) " in Least_in_Ord, (assumption+))
 apply (rule DiffI, best intro: Ord_in_Ord LeastI, clarify)
-apply (erule_tac P = "%z. x \<in> F (z) " and i = c in less_LeastE)
+apply (erule_tac P = "\<lambda>z. x \<in> F (z) " and i = c in less_LeastE)
 apply (blast intro: Ord_Least ltI)
 done
 
--- a/src/ZF/AC/DC.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/DC.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -11,7 +11,7 @@
 lemma RepFun_lepoll: "Ord(a) \<Longrightarrow> {P(b). b \<in> a} \<lesssim> a"
 apply (unfold lepoll_def)
 apply (rule_tac x = "\<lambda>z \<in> RepFun (a,P) . \<mu> i. z=P (i) " in exI)
-apply (rule_tac d="%z. P (z)" in lam_injective)
+apply (rule_tac d="\<lambda>z. P (z)" in lam_injective)
  apply (fast intro!: Least_in_Ord)
 apply (rule sym) 
 apply (fast intro: LeastI Ord_in_Ord) 
@@ -21,7 +21,7 @@
 lemma image_Ord_lepoll: "\<lbrakk>f \<in> X->Y; Ord(X)\<rbrakk> \<Longrightarrow> f``X \<lesssim> X"
 apply (unfold lepoll_def)
 apply (rule_tac x = "\<lambda>x \<in> f``X. \<mu> y. f`y = x" in exI)
-apply (rule_tac d = "%z. f`z" in lam_injective)
+apply (rule_tac d = "\<lambda>z. f`z" in lam_injective)
 apply (fast intro!: Least_in_Ord apply_equality, clarify) 
 apply (rule LeastI) 
  apply (erule apply_equality, assumption+) 
@@ -29,35 +29,35 @@
 done
 
 lemma range_subset_domain: 
-      "\<lbrakk>R \<subseteq> X*X;   \<And>g. g \<in> X \<Longrightarrow> \<exists>u. <g,u> \<in> R\<rbrakk> 
+      "\<lbrakk>R \<subseteq> X*X;   \<And>g. g \<in> X \<Longrightarrow> \<exists>u. \<langle>g,u\<rangle> \<in> R\<rbrakk> 
        \<Longrightarrow> range(R) \<subseteq> domain(R)"
 by blast 
 
-lemma cons_fun_type: "g \<in> n->X \<Longrightarrow> cons(<n,x>, g) \<in> succ(n) -> cons(x, X)"
+lemma cons_fun_type: "g \<in> n->X \<Longrightarrow> cons(\<langle>n,x\<rangle>, g) \<in> succ(n) -> cons(x, X)"
 apply (unfold succ_def)
 apply (fast intro!: fun_extend elim!: mem_irrefl)
 done
 
 lemma cons_fun_type2:
-     "\<lbrakk>g \<in> n->X; x \<in> X\<rbrakk> \<Longrightarrow> cons(<n,x>, g) \<in> succ(n) -> X"
+     "\<lbrakk>g \<in> n->X; x \<in> X\<rbrakk> \<Longrightarrow> cons(\<langle>n,x\<rangle>, g) \<in> succ(n) -> X"
 by (erule cons_absorb [THEN subst], erule cons_fun_type)
 
-lemma cons_image_n: "n \<in> nat \<Longrightarrow> cons(<n,x>, g)``n = g``n"
+lemma cons_image_n: "n \<in> nat \<Longrightarrow> cons(\<langle>n,x\<rangle>, g)``n = g``n"
 by (fast elim!: mem_irrefl)
 
-lemma cons_val_n: "g \<in> n->X \<Longrightarrow> cons(<n,x>, g)`n = x"
+lemma cons_val_n: "g \<in> n->X \<Longrightarrow> cons(\<langle>n,x\<rangle>, g)`n = x"
 by (fast intro!: apply_equality elim!: cons_fun_type)
 
-lemma cons_image_k: "k \<in> n \<Longrightarrow> cons(<n,x>, g)``k = g``k"
+lemma cons_image_k: "k \<in> n \<Longrightarrow> cons(\<langle>n,x\<rangle>, g)``k = g``k"
 by (fast elim: mem_asym)
 
-lemma cons_val_k: "\<lbrakk>k \<in> n; g \<in> n->X\<rbrakk> \<Longrightarrow> cons(<n,x>, g)`k = g`k"
+lemma cons_val_k: "\<lbrakk>k \<in> n; g \<in> n->X\<rbrakk> \<Longrightarrow> cons(\<langle>n,x\<rangle>, g)`k = g`k"
 by (fast intro!: apply_equality consI2 elim!: cons_fun_type apply_Pair)
 
-lemma domain_cons_eq_succ: "domain(f)=x \<Longrightarrow> domain(cons(<x,y>, f)) = succ(x)"
+lemma domain_cons_eq_succ: "domain(f)=x \<Longrightarrow> domain(cons(\<langle>x,y\<rangle>, f)) = succ(x)"
 by (simp add: domain_cons succ_def)
 
-lemma restrict_cons_eq: "g \<in> n->X \<Longrightarrow> restrict(cons(<n,x>, g), n) = g"
+lemma restrict_cons_eq: "g \<in> n->X \<Longrightarrow> restrict(cons(\<langle>n,x\<rangle>, g), n) = g"
 apply (simp add: restrict_def Pi_iff)
 apply (blast intro: elim: mem_irrefl)  
 done
@@ -80,9 +80,9 @@
 
 
 definition
-  DC  :: "i => o"  where
+  DC  :: "i \<Rightarrow> o"  where
     "DC(a) \<equiv> \<forall>X R. R \<subseteq> Pow(X)*X  \<and>
-                    (\<forall>Y \<in> Pow(X). Y \<prec> a \<longrightarrow> (\<exists>x \<in> X. <Y,x> \<in> R)) 
+                    (\<forall>Y \<in> Pow(X). Y \<prec> a \<longrightarrow> (\<exists>x \<in> X. \<langle>Y,x\<rangle> \<in> R)) 
                     \<longrightarrow> (\<exists>f \<in> a->X. \<forall>b<a. <f``b,f`b> \<in> R)"
 
 definition
@@ -91,18 +91,18 @@
                     \<longrightarrow> (\<exists>f \<in> nat->domain(R). \<forall>n \<in> nat. <f`n,f`succ(n)>:R)"
 
 definition
-  ff  :: "[i, i, i, i] => i"  where
+  ff  :: "[i, i, i, i] \<Rightarrow> i"  where
     "ff(b, X, Q, R) \<equiv>
-           transrec(b, %c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))"
+           transrec(b, \<lambda>c r. THE x. first(x, {x \<in> X. <r``c, x> \<in> R}, Q))"
 
 
 locale DC0_imp =
   fixes XX and RR and X and R
 
-  assumes all_ex: "\<forall>Y \<in> Pow(X). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R)"
+  assumes all_ex: "\<forall>Y \<in> Pow(X). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> X. \<langle>Y, x\<rangle> \<in> R)"
 
   defines XX_def: "XX \<equiv> (\<Union>n \<in> nat. {f \<in> n->X. \<forall>k \<in> n. <f``k, f`k> \<in> R})"
-     and RR_def:  "RR \<equiv> {<z1,z2>:XX*XX. domain(z2)=succ(domain(z1))  
+     and RR_def:  "RR \<equiv> {\<langle>z1,z2\<rangle>:XX*XX. domain(z2)=succ(domain(z1))  
                                        \<and> restrict(z2, domain(z1)) = z1}"
 begin
 
@@ -141,7 +141,7 @@
 apply (erule_tac [2] notE [OF _ empty_subsetI [THEN PowI]])
 apply (erule_tac impE [OF _ nat_0I [THEN n_lesspoll_nat]])
 apply (erule bexE)
-apply (rule_tac a = "<0, {<0, x>}>" in not_emptyI)
+apply (rule_tac a = "<0, {\<langle>0, x\<rangle>}>" in not_emptyI)
 apply (rule CollectI)
 apply (rule SigmaI)
 apply (rule nat_0I [THEN UN_I])
@@ -161,7 +161,7 @@
                                           [OF _ nat_into_Ord] n_lesspoll_nat]],
        assumption+)
 apply (erule bexE)
-apply (rule_tac x = "cons (<n,x>, g) " in exI)
+apply (rule_tac x = "cons (\<langle>n,x\<rangle>, g) " in exI)
 apply (rule CollectI)
 apply (force elim!: cons_fun_type2 
              simp add: cons_image_n cons_val_n cons_image_k cons_val_k)
@@ -266,12 +266,12 @@
                                                                           
     XX = (\<Union>n \<in> nat. {f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})
 
-    RR = {<z1,z2>:Fin(XX)*XX. 
+    RR = {\<langle>z1,z2\<rangle>:Fin(XX)*XX. 
            (domain(z2)=succ(\<Union>f \<in> z1. domain(f)) \<and>
             (\<forall>f \<in> z1. restrict(z2, domain(f)) = f)) |      
            (\<not> (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f)) \<and>
                         (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) \<and>           
-            z2={<0,x>})}                                          
+            z2={\<langle>0,x\<rangle>})}                                          
                                                                           
    Then XX and RR satisfy the hypotheses of DC(omega).                    
    So applying DC:                                                        
@@ -287,7 +287,7 @@
 ************************************************************************* *)
 
 lemma singleton_in_funs: 
- "x \<in> X \<Longrightarrow> {<0,x>} \<in> 
+ "x \<in> X \<Longrightarrow> {\<langle>0,x\<rangle>} \<in> 
             (\<Union>n \<in> nat. {f \<in> succ(n)->X. \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
 apply (rule nat_0I [THEN UN_I])
 apply (force simp add: singleton_0 [symmetric]
@@ -300,15 +300,15 @@
   defines XX_def: "XX \<equiv> (\<Union>n \<in> nat.
                       {f \<in> succ(n)->domain(R). \<forall>k \<in> n. <f`k, f`succ(k)> \<in> R})"
       and RR_def:
-         "RR \<equiv> {<z1,z2>:Fin(XX)*XX. 
+         "RR \<equiv> {\<langle>z1,z2\<rangle>:Fin(XX)*XX. 
                   (domain(z2)=succ(\<Union>f \<in> z1. domain(f))  
                     \<and> (\<forall>f \<in> z1. restrict(z2, domain(f)) = f))
                   | (\<not> (\<exists>g \<in> XX. domain(g)=succ(\<Union>f \<in> z1. domain(f))  
-                     \<and> (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) \<and> z2={<0,x>})}"
+                     \<and> (\<forall>f \<in> z1. restrict(g, domain(f)) = f)) \<and> z2={\<langle>0,x\<rangle>})}"
       and allRR_def:
         "allRR \<equiv> \<forall>b<nat.
                    <f``b, f`b> \<in>  
-                    {<z1,z2>\<in>Fin(XX)*XX. (domain(z2)=succ(\<Union>f \<in> z1. domain(f))
+                    {\<langle>z1,z2\<rangle>\<in>Fin(XX)*XX. (domain(z2)=succ(\<Union>f \<in> z1. domain(f))
                                     \<and> (\<Union>f \<in> z1. domain(f)) = b  
                                     \<and> (\<forall>f \<in> z1. restrict(z2,domain(f)) = f))}"
 begin
@@ -316,7 +316,7 @@
 lemma lemma4:
      "\<lbrakk>range(R) \<subseteq> domain(R);  x \<in> domain(R)\<rbrakk>   
       \<Longrightarrow> RR \<subseteq> Pow(XX)*XX \<and>   
-             (\<forall>Y \<in> Pow(XX). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> XX. <Y,x>:RR))"
+             (\<forall>Y \<in> Pow(XX). Y \<prec> nat \<longrightarrow> (\<exists>x \<in> XX. \<langle>Y,x\<rangle>:RR))"
 apply (rule conjI)
 apply (force dest!: FinD [THEN PowI] simp add: RR_def)
 apply (rule impI [THEN ballI])
@@ -368,7 +368,7 @@
 
 lemma restrict_cons_eq_restrict: 
      "\<lbrakk>restrict(h, domain(u))=u;  h \<in> n->X;  domain(u) \<subseteq> n\<rbrakk>   
-      \<Longrightarrow> restrict(cons(<n, y>, h), domain(u)) = u"
+      \<Longrightarrow> restrict(cons(\<langle>n, y\<rangle>, h), domain(u)) = u"
 apply (unfold restrict_def)
 apply (simp add: restrict_def Pi_iff)
 apply (erule sym [THEN trans, symmetric])
@@ -512,7 +512,7 @@
 apply (erule allE impE)+
 apply (rule Card_Hartog)
 apply (erule_tac x = A in allE)
-apply (erule_tac x = "{<z1,z2> \<in> Pow (A) *A . z1 \<prec> Hartog (A) \<and> z2 \<notin> z1}" 
+apply (erule_tac x = "{\<langle>z1,z2\<rangle> \<in> Pow (A) *A . z1 \<prec> Hartog (A) \<and> z2 \<notin> z1}" 
                  in allE)
 apply simp
 apply (erule impE, fast elim: lesspoll_lemma [THEN not_emptyE])
@@ -549,15 +549,15 @@
 by (fast intro!: lam_type RepFunI)
 
 lemma lemmaX:
-     "\<lbrakk>\<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R);   
+     "\<lbrakk>\<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. \<langle>Y, x\<rangle> \<in> R);   
          b \<in> K; Z \<in> Pow(X); Z \<prec> K\<rbrakk>   
-      \<Longrightarrow> {x \<in> X. <Z,x> \<in> R} \<noteq> 0"
+      \<Longrightarrow> {x \<in> X. \<langle>Z,x\<rangle> \<in> R} \<noteq> 0"
 by blast
 
 
 lemma WO1_DC_lemma:
      "\<lbrakk>Card(K); well_ord(X,Q);   
-         \<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. <Y, x> \<in> R); b \<in> K\<rbrakk>   
+         \<forall>Y \<in> Pow(X). Y \<prec> K \<longrightarrow> (\<exists>x \<in> X. \<langle>Y, x\<rangle> \<in> R); b \<in> K\<rbrakk>   
       \<Longrightarrow> ff(b, X, Q, R) \<in> {x \<in> X. <(\<lambda>c \<in> b. ff(c, X, Q, R))``b, x> \<in> R}"
 apply (rule_tac P = "b \<in> K" in impE, (erule_tac [2] asm_rl)+)
 apply (rule_tac i=b in Card_is_Ord [THEN Ord_in_Ord, THEN trans_induct], 
--- a/src/ZF/AC/HH.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/HH.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,8 +12,8 @@
 begin
 
 definition
-  HH :: "[i, i, i] => i"  where
-    "HH(f,x,a) \<equiv> transrec(a, %b r. let z = x - (\<Union>c \<in> b. r`c)
+  HH :: "[i, i, i] \<Rightarrow> i"  where
+    "HH(f,x,a) \<equiv> transrec(a, \<lambda>b r. let z = x - (\<Union>c \<in> b. r`c)
                                     in  if f`z \<in> Pow(z)-{0} then f`z else {x})"
 
 subsection\<open>Lemmas useful in each of the three proofs\<close>
@@ -57,7 +57,7 @@
 prefer 2 apply assumption+
 apply (rule leI [THEN le_imp_subset, THEN subset_imp_Diff_eq, THEN ssubst], 
        assumption)
-apply (rule_tac t = "%z. z-X" for X in subst_context)
+apply (rule_tac t = "\<lambda>z. z-X" for X in subst_context)
 apply (rule Diff_UN_eq_self)
 apply (drule Ord_DiffE, assumption) 
 apply (fast elim: ltE, auto) 
@@ -82,7 +82,7 @@
 
 lemma HH_eq_arg_lt:
      "\<lbrakk>HH(f,x,v)=HH(f,x,w); HH(f,x,v) \<in> Pow(x)-{0}; v \<in> w\<rbrakk> \<Longrightarrow> P"
-apply (frule_tac P = "%y. y \<in> Pow (x) -{0}" in subst, assumption)
+apply (frule_tac P = "\<lambda>y. y \<in> Pow (x) -{0}" in subst, assumption)
 apply (drule_tac a = w in HH_subset_x_imp_subset_Diff_UN)
 apply (drule subst_elem, assumption)
 apply (fast intro!: singleton_iff [THEN iffD2] equals0I)
--- a/src/ZF/AC/Hartog.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/Hartog.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,7 +9,7 @@
 begin
 
 definition
-  Hartog :: "i => i"  where
+  Hartog :: "i \<Rightarrow> i"  where
    "Hartog(X) \<equiv> \<mu> i. \<not> i \<lesssim> X"
 
 lemma Ords_in_set: "\<forall>a. Ord(a) \<longrightarrow> a \<in> X \<Longrightarrow> P"
@@ -45,7 +45,7 @@
 lemma Ords_lepoll_set_lemma:
      "(\<forall>a. Ord(a) \<longrightarrow> a \<lesssim> X) \<Longrightarrow>   
        \<forall>a. Ord(a) \<longrightarrow>   
-        a \<in> {b. Z \<in> Pow(X)*Pow(X*X), \<exists>Y R. Z=<Y,R> \<and> ordertype(Y,R)=b}"
+        a \<in> {b. Z \<in> Pow(X)*Pow(X*X), \<exists>Y R. Z=\<langle>Y,R\<rangle> \<and> ordertype(Y,R)=b}"
 apply (intro allI impI)
 apply (elim allE impE, assumption)
 apply (blast dest!: Ord_lepoll_imp_eq_ordertype intro: sym) 
--- a/src/ZF/AC/WO2_AC16.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/WO2_AC16.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -18,10 +18,10 @@
 (**** A recursive definition used in the proof of WO2 \<Longrightarrow> AC16 ****)
 
 definition
-  recfunAC16 :: "[i,i,i,i] => i"  where
+  recfunAC16 :: "[i,i,i,i] \<Rightarrow> i"  where
     "recfunAC16(f,h,i,a) \<equiv> 
          transrec2(i, 0, 
-              %g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r
+              \<lambda>g r. if (\<exists>y \<in> r. h`g \<subseteq> y) then r
                     else r \<union> {f`(\<mu> i. h`g \<subseteq> f`i \<and> 
                          (\<forall>b<a. (h`b \<subseteq> f`i \<longrightarrow> (\<forall>t \<in> r. \<not> h`b \<subseteq> t))))})"
 
@@ -224,7 +224,7 @@
                           empty_subsetI [THEN subset_imp_lepoll])
 (*succ case*)
 apply (simp add: recfunAC16_succ
-                 Diff_UN_succ_empty [of _ "%j. recfunAC16(f,g,j,a)"]
+                 Diff_UN_succ_empty [of _ "\<lambda>j. recfunAC16(f,g,j,a)"]
                  empty_subsetI [THEN subset_imp_lepoll])
 apply (best intro: Diff_UN_succ_subset [THEN subset_imp_lepoll]
                    singleton_eqpoll_1 [THEN eqpoll_imp_lepoll] lepoll_trans)
@@ -563,7 +563,7 @@
 apply (elim exE)
 apply (rename_tac h)
 apply (rule_tac x = "\<Union>j<a. recfunAC16 (h,f,j,a) " in exI)
-apply (rule_tac P="%z. Y \<and> (\<forall>x \<in> z. Z(x))"  for Y Z
+apply (rule_tac P="\<lambda>z. Y \<and> (\<forall>x \<in> z. Z(x))"  for Y Z
        in bij_is_surj [THEN surj_image_eq, THEN subst], assumption)
 apply (rule lemma_simp_induct)
 apply (blast del: conjI notI
--- a/src/ZF/AC/WO6_WO1.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/AC/WO6_WO1.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -16,18 +16,18 @@
 
 (* Auxiliary definitions used in proof *)
 definition
-  NN  :: "i => i"  where
+  NN  :: "i \<Rightarrow> i"  where
     "NN(y) \<equiv> {m \<in> nat. \<exists>a. \<exists>f. Ord(a) \<and> domain(f)=a  \<and>  
                         (\<Union>b<a. f`b) = y \<and> (\<forall>b<a. f`b \<lesssim> m)}"
   
 definition
-  uu  :: "[i, i, i, i] => i"  where
+  uu  :: "[i, i, i, i] \<Rightarrow> i"  where
     "uu(f, beta, gamma, delta) \<equiv> (f`beta * f`gamma) \<inter> f`delta"
 
 
 (** Definitions for case 1 **)
 definition
-  vv1 :: "[i, i, i] => i"  where
+  vv1 :: "[i, i, i] \<Rightarrow> i"  where
      "vv1(f,m,b) \<equiv>                                                
            let g = \<mu> g. (\<exists>d. Ord(d) \<and> (domain(uu(f,b,g,d)) \<noteq> 0 \<and> 
                                  domain(uu(f,b,g,d)) \<lesssim> m));      
@@ -36,26 +36,26 @@
            in  if f`b \<noteq> 0 then domain(uu(f,b,g,d)) else 0"
 
 definition
-  ww1 :: "[i, i, i] => i"  where
+  ww1 :: "[i, i, i] \<Rightarrow> i"  where
      "ww1(f,m,b) \<equiv> f`b - vv1(f,m,b)"
 
 definition
-  gg1 :: "[i, i, i] => i"  where
+  gg1 :: "[i, i, i] \<Rightarrow> i"  where
      "gg1(f,a,m) \<equiv> \<lambda>b \<in> a++a. if b<a then vv1(f,m,b) else ww1(f,m,b--a)"
 
 
 (** Definitions for case 2 **)
 definition
-  vv2 :: "[i, i, i, i] => i"  where
+  vv2 :: "[i, i, i, i] \<Rightarrow> i"  where
      "vv2(f,b,g,s) \<equiv>   
            if f`g \<noteq> 0 then {uu(f, b, g, \<mu> d. uu(f,b,g,d) \<noteq> 0)`s} else 0"
 
 definition
-  ww2 :: "[i, i, i, i] => i"  where
+  ww2 :: "[i, i, i, i] \<Rightarrow> i"  where
      "ww2(f,b,g,s) \<equiv> f`g - vv2(f,b,g,s)"
 
 definition
-  gg2 :: "[i, i, i, i] => i"  where
+  gg2 :: "[i, i, i, i] \<Rightarrow> i"  where
      "gg2(f,a,b,s) \<equiv>
               \<lambda>g \<in> a++a. if g<a then vv2(f,b,g,s) else ww2(f,b,g--a,s)"
 
@@ -213,12 +213,12 @@
          Least_a = (\<mu> a. \<exists>x. Ord(x) \<and> P(a, x))\<rbrakk>   
       \<Longrightarrow> P(Least_a, \<mu> b. P(Least_a, b))"
 apply (erule ssubst)
-apply (rule_tac Q = "%z. P (z, \<mu> b. P (z, b))" in LeastI2)
+apply (rule_tac Q = "\<lambda>z. P (z, \<mu> b. P (z, b))" in LeastI2)
 apply (fast elim!: LeastI)+
 done
 
 lemmas nested_Least_instance =
-       nested_LeastI [of "%g d. domain(uu(f,b,g,d)) \<noteq> 0 \<and> 
+       nested_LeastI [of "\<lambda>g d. domain(uu(f,b,g,d)) \<noteq> 0 \<and> 
                                 domain(uu(f,b,g,d)) \<lesssim> m"] for f b m
 
 lemma gg1_lepoll_m: 
@@ -409,26 +409,26 @@
    (used only in the following two lemmas)                          *)
 
 lemma z_n_subset_z_succ_n:
-     "\<forall>n \<in> nat. rec(n, x, %k r. r \<union> r*r) \<subseteq> rec(succ(n), x, %k r. r \<union> r*r)"
+     "\<forall>n \<in> nat. rec(n, x, \<lambda>k r. r \<union> r*r) \<subseteq> rec(succ(n), x, \<lambda>k r. r \<union> r*r)"
 by (fast intro: rec_succ [THEN ssubst])
 
 lemma le_subsets:
      "\<lbrakk>\<forall>n \<in> nat. f(n)<=f(succ(n)); n\<le>m; n \<in> nat; m \<in> nat\<rbrakk>   
       \<Longrightarrow> f(n)<=f(m)"
 apply (erule_tac P = "n\<le>m" in rev_mp)
-apply (rule_tac P = "%z. n\<le>z \<longrightarrow> f (n) \<subseteq> f (z) " in nat_induct) 
+apply (rule_tac P = "\<lambda>z. n\<le>z \<longrightarrow> f (n) \<subseteq> f (z) " in nat_induct) 
 apply (auto simp add: le_iff) 
 done
 
 lemma le_imp_rec_subset:
      "\<lbrakk>n\<le>m; m \<in> nat\<rbrakk> 
-      \<Longrightarrow> rec(n, x, %k r. r \<union> r*r) \<subseteq> rec(m, x, %k r. r \<union> r*r)"
+      \<Longrightarrow> rec(n, x, \<lambda>k r. r \<union> r*r) \<subseteq> rec(m, x, \<lambda>k r. r \<union> r*r)"
 apply (rule z_n_subset_z_succ_n [THEN le_subsets])
 apply (blast intro: lt_nat_in_nat)+
 done
 
 lemma lemma_iv: "\<exists>y. x \<union> y*y \<subseteq> y"
-apply (rule_tac x = "\<Union>n \<in> nat. rec (n, x, %k r. r \<union> r*r) " in exI)
+apply (rule_tac x = "\<Union>n \<in> nat. rec (n, x, \<lambda>k r. r \<union> r*r) " in exI)
 apply safe
 apply (rule nat_0I [THEN UN_I], simp)
 apply (rule_tac a = "succ (n \<union> na) " in UN_I)
@@ -477,7 +477,7 @@
 apply (rule_tac x = a in exI)
 apply (rule_tac x = "\<lambda>x \<in> y. \<mu> i. f`i = {x}" in exI)
 apply (rule conjI, assumption)
-apply (rule_tac d = "%i. THE x. x \<in> f`i" in lam_injective)
+apply (rule_tac d = "\<lambda>i. THE x. x \<in> f`i" in lam_injective)
 apply (drule lemma1, assumption+)
 apply (fast elim!: Least_le [THEN lt_trans1, THEN ltD] lt_Ord)
 apply (rule lemma2 [THEN ssubst], assumption+, blast)
--- a/src/ZF/Arith.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Arith.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -6,7 +6,7 @@
 (*"Difference" is subtraction of natural numbers.
   There are no negative numbers; we have
      m #- n = 0  iff  m<=n   and     m #- n = succ(k) iff m>n.
-  Also, rec(m, 0, %z w.z) is pred(m).
+  Also, rec(m, 0, \<lambda>z w.z) is pred(m).
 *)
 
 section\<open>Arithmetic Operators and Their Definitions\<close>
@@ -16,18 +16,18 @@
 text\<open>Proofs about elementary arithmetic: addition, multiplication, etc.\<close>
 
 definition
-  pred   :: "i=>i"    (*inverse of succ*)  where
-    "pred(y) \<equiv> nat_case(0, %x. x, y)"
+  pred   :: "i\<Rightarrow>i"    (*inverse of succ*)  where
+    "pred(y) \<equiv> nat_case(0, \<lambda>x. x, y)"
 
 definition
-  natify :: "i=>i"    (*coerces non-nats to nats*)  where
-    "natify \<equiv> Vrecursor(%f a. if a = succ(pred(a)) then succ(f`pred(a))
+  natify :: "i\<Rightarrow>i"    (*coerces non-nats to nats*)  where
+    "natify \<equiv> Vrecursor(\<lambda>f a. if a = succ(pred(a)) then succ(f`pred(a))
                                                     else 0)"
 
 consts
-  raw_add  :: "[i,i]=>i"
-  raw_diff  :: "[i,i]=>i"
-  raw_mult  :: "[i,i]=>i"
+  raw_add  :: "[i,i]\<Rightarrow>i"
+  raw_diff  :: "[i,i]\<Rightarrow>i"
+  raw_mult  :: "[i,i]\<Rightarrow>i"
 
 primrec
   "raw_add (0, n) = n"
@@ -36,40 +36,40 @@
 primrec
   raw_diff_0:     "raw_diff(m, 0) = m"
   raw_diff_succ:  "raw_diff(m, succ(n)) =
-                     nat_case(0, %x. x, raw_diff(m, n))"
+                     nat_case(0, \<lambda>x. x, raw_diff(m, n))"
 
 primrec
   "raw_mult(0, n) = 0"
   "raw_mult(succ(m), n) = raw_add (n, raw_mult(m, n))"
 
 definition
-  add :: "[i,i]=>i"                    (infixl \<open>#+\<close> 65)  where
+  add :: "[i,i]\<Rightarrow>i"                    (infixl \<open>#+\<close> 65)  where
     "m #+ n \<equiv> raw_add (natify(m), natify(n))"
 
 definition
-  diff :: "[i,i]=>i"                    (infixl \<open>#-\<close> 65)  where
+  diff :: "[i,i]\<Rightarrow>i"                    (infixl \<open>#-\<close> 65)  where
     "m #- n \<equiv> raw_diff (natify(m), natify(n))"
 
 definition
-  mult :: "[i,i]=>i"                    (infixl \<open>#*\<close> 70)  where
+  mult :: "[i,i]\<Rightarrow>i"                    (infixl \<open>#*\<close> 70)  where
     "m #* n \<equiv> raw_mult (natify(m), natify(n))"
 
 definition
-  raw_div  :: "[i,i]=>i"  where
+  raw_div  :: "[i,i]\<Rightarrow>i"  where
     "raw_div (m, n) \<equiv>
-       transrec(m, %j f. if j<n | n=0 then 0 else succ(f`(j#-n)))"
+       transrec(m, \<lambda>j f. if j<n | n=0 then 0 else succ(f`(j#-n)))"
 
 definition
-  raw_mod  :: "[i,i]=>i"  where
+  raw_mod  :: "[i,i]\<Rightarrow>i"  where
     "raw_mod (m, n) \<equiv>
-       transrec(m, %j f. if j<n | n=0 then j else f`(j#-n))"
+       transrec(m, \<lambda>j f. if j<n | n=0 then j else f`(j#-n))"
 
 definition
-  div  :: "[i,i]=>i"                    (infixl \<open>div\<close> 70)  where
+  div  :: "[i,i]\<Rightarrow>i"                    (infixl \<open>div\<close> 70)  where
     "m div n \<equiv> raw_div (natify(m), natify(n))"
 
 definition
-  mod  :: "[i,i]=>i"                    (infixl \<open>mod\<close> 70)  where
+  mod  :: "[i,i]\<Rightarrow>i"                    (infixl \<open>mod\<close> 70)  where
     "m mod n \<equiv> raw_mod (natify(m), natify(n))"
 
 declare rec_type [simp]
@@ -337,7 +337,7 @@
 
 text\<open>\<open>\<le>\<close> monotonicity, 1st argument\<close>
 lemma add_le_mono1: "\<lbrakk>i \<le> j; j\<in>nat\<rbrakk> \<Longrightarrow> i#+k \<le> j#+k"
-apply (rule_tac f = "%j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck)
+apply (rule_tac f = "\<lambda>j. j#+k" in Ord_lt_mono_imp_le_mono, typecheck)
 apply (blast intro: add_lt_mono1 add_type [THEN nat_into_Ord])+
 done
 
--- a/src/ZF/Bin.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Bin.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -26,13 +26,13 @@
         | Bit ("w \<in> bin", "b \<in> bool")     (infixl \<open>BIT\<close> 90)
 
 consts
-  integ_of  :: "i=>i"
-  NCons     :: "[i,i]=>i"
-  bin_succ  :: "i=>i"
-  bin_pred  :: "i=>i"
-  bin_minus :: "i=>i"
-  bin_adder :: "i=>i"
-  bin_mult  :: "[i,i]=>i"
+  integ_of  :: "i\<Rightarrow>i"
+  NCons     :: "[i,i]\<Rightarrow>i"
+  bin_succ  :: "i\<Rightarrow>i"
+  bin_pred  :: "i\<Rightarrow>i"
+  bin_minus :: "i\<Rightarrow>i"
+  bin_adder :: "i\<Rightarrow>i"
+  bin_mult  :: "[i,i]\<Rightarrow>i"
 
 primrec
   integ_of_Pls:  "integ_of (Pls)     = $# 0"
@@ -74,7 +74,7 @@
     "bin_adder (v BIT x) =
        (\<lambda>w\<in>bin.
          bin_case (v BIT x, bin_pred(v BIT x),
-                   %w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w),
+                   \<lambda>w y. NCons(bin_adder (v) ` cond(x and y, bin_succ(w), w),
                                x xor y),
                    w))"
 
@@ -87,7 +87,7 @@
 *)
 
 definition
-  bin_add   :: "[i,i]=>i"  where
+  bin_add   :: "[i,i]\<Rightarrow>i"  where
     "bin_add(v,w) \<equiv> bin_adder(v)`w"
 
 
@@ -114,8 +114,8 @@
   "#-2" \<rightleftharpoons> "CONST integ_of(CONST Min BIT 0)"
 
 syntax
-  "_Int" :: "num_token => i"  (\<open>#_\<close> 1000)
-  "_Neg_Int" :: "num_token => i"  (\<open>#-_\<close> 1000)
+  "_Int" :: "num_token \<Rightarrow> i"  (\<open>#_\<close> 1000)
+  "_Neg_Int" :: "num_token \<Rightarrow> i"  (\<open>#-_\<close> 1000)
 
 ML_file \<open>Tools/numeral_syntax.ML\<close>
 
--- a/src/ZF/Bool.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Bool.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -24,15 +24,15 @@
 definition "not(b) \<equiv> cond(b,0,1)"
 
 definition
-  "and"       :: "[i,i]=>i"      (infixl \<open>and\<close> 70)  where
+  "and"       :: "[i,i]\<Rightarrow>i"      (infixl \<open>and\<close> 70)  where
     "a and b \<equiv> cond(a,b,0)"
 
 definition
-  or          :: "[i,i]=>i"      (infixl \<open>or\<close> 65)  where
+  or          :: "[i,i]\<Rightarrow>i"      (infixl \<open>or\<close> 65)  where
     "a or b \<equiv> cond(a,1,b)"
 
 definition
-  xor         :: "[i,i]=>i"      (infixl \<open>xor\<close> 65) where
+  xor         :: "[i,i]\<Rightarrow>i"      (infixl \<open>xor\<close> 65) where
     "a xor b \<equiv> cond(a,not(b),b)"
 
 
@@ -152,7 +152,7 @@
 
 
 definition
-  bool_of_o :: "o=>i" where
+  bool_of_o :: "o\<Rightarrow>i" where
    "bool_of_o(P) \<equiv> (if P then 1 else 0)"
 
 lemma [simp]: "bool_of_o(True) = 1"
--- a/src/ZF/Cardinal.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Cardinal.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,31 +9,31 @@
 
 definition
   (*least ordinal operator*)
-   Least    :: "(i=>o) => i"    (binder \<open>\<mu> \<close> 10)  where
+   Least    :: "(i\<Rightarrow>o) \<Rightarrow> i"    (binder \<open>\<mu> \<close> 10)  where
      "Least(P) \<equiv> THE i. Ord(i) \<and> P(i) \<and> (\<forall>j. j<i \<longrightarrow> \<not>P(j))"
 
 definition
-  eqpoll   :: "[i,i] => o"     (infixl \<open>\<approx>\<close> 50)  where
+  eqpoll   :: "[i,i] \<Rightarrow> o"     (infixl \<open>\<approx>\<close> 50)  where
     "A \<approx> B \<equiv> \<exists>f. f \<in> bij(A,B)"
 
 definition
-  lepoll   :: "[i,i] => o"     (infixl \<open>\<lesssim>\<close> 50)  where
+  lepoll   :: "[i,i] \<Rightarrow> o"     (infixl \<open>\<lesssim>\<close> 50)  where
     "A \<lesssim> B \<equiv> \<exists>f. f \<in> inj(A,B)"
 
 definition
-  lesspoll :: "[i,i] => o"     (infixl \<open>\<prec>\<close> 50)  where
+  lesspoll :: "[i,i] \<Rightarrow> o"     (infixl \<open>\<prec>\<close> 50)  where
     "A \<prec> B \<equiv> A \<lesssim> B \<and> \<not>(A \<approx> B)"
 
 definition
-  cardinal :: "i=>i"           (\<open>|_|\<close>)  where
+  cardinal :: "i\<Rightarrow>i"           (\<open>|_|\<close>)  where
     "|A| \<equiv> (\<mu> i. i \<approx> A)"
 
 definition
-  Finite   :: "i=>o"  where
+  Finite   :: "i\<Rightarrow>o"  where
     "Finite(A) \<equiv> \<exists>n\<in>nat. A \<approx> n"
 
 definition
-  Card     :: "i=>o"  where
+  Card     :: "i\<Rightarrow>o"  where
     "Card(i) \<equiv> (i = |i|)"
 
 
@@ -42,14 +42,14 @@
 
 (** Lemma: Banach's Decomposition Theorem **)
 
-lemma decomp_bnd_mono: "bnd_mono(X, %W. X - g``(Y - f``W))"
+lemma decomp_bnd_mono: "bnd_mono(X, \<lambda>W. X - g``(Y - f``W))"
 by (rule bnd_monoI, blast+)
 
 lemma Banach_last_equation:
     "g \<in> Y->X
-     \<Longrightarrow> g``(Y - f`` lfp(X, %W. X - g``(Y - f``W))) =
-         X - lfp(X, %W. X - g``(Y - f``W))"
-apply (rule_tac P = "%u. v = X-u" for v
+     \<Longrightarrow> g``(Y - f`` lfp(X, \<lambda>W. X - g``(Y - f``W))) =
+         X - lfp(X, \<lambda>W. X - g``(Y - f``W))"
+apply (rule_tac P = "\<lambda>u. v = X-u" for v
        in decomp_bnd_mono [THEN lfp_unfold, THEN ssubst])
 apply (simp add: double_complement  fun_is_rel [THEN image_subset])
 done
@@ -738,7 +738,7 @@
     "\<lbrakk>A \<lesssim> B;  b \<notin> B\<rbrakk> \<Longrightarrow> cons(a,A) \<lesssim> cons(b,B)"
 apply (unfold lepoll_def, safe)
 apply (rule_tac x = "\<lambda>y\<in>cons (a,A) . if y=a then b else f`y" in exI)
-apply (rule_tac d = "%z. if z \<in> B then converse (f) `z else a" in lam_injective)
+apply (rule_tac d = "\<lambda>z. if z \<in> B then converse (f) `z else a" in lam_injective)
 apply (safe elim!: consE')
    apply simp_all
 apply (blast intro: inj_is_fun [THEN apply_type])+
@@ -769,7 +769,7 @@
 lemma not_0_is_lepoll_1: "A \<noteq> 0 \<Longrightarrow> 1 \<lesssim> A"
 apply (erule not_emptyE)
 apply (rule_tac a = "cons (x, A-{x}) " in subst)
-apply (rule_tac [2] a = "cons(0,0)" and P= "%y. y \<lesssim> cons (x, A-{x})" in subst)
+apply (rule_tac [2] a = "cons(0,0)" and P= "\<lambda>y. y \<lesssim> cons (x, A-{x})" in subst)
 prefer 3 apply (blast intro: cons_lepoll_cong subset_imp_lepoll, auto)
 done
 
@@ -796,8 +796,8 @@
     "\<lbrakk>f \<in> inj(A,B);  A \<inter> B = 0\<rbrakk> \<Longrightarrow> A \<union> (B - range(f)) \<approx> B"
 apply (unfold eqpoll_def)
 apply (rule exI)
-apply (rule_tac c = "%x. if x \<in> A then f`x else x"
-            and d = "%y. if y \<in> range (f) then converse (f) `y else y"
+apply (rule_tac c = "\<lambda>x. if x \<in> A then f`x else x"
+            and d = "\<lambda>y. if y \<in> range (f) then converse (f) `y else y"
        in lam_bijective)
 apply (blast intro!: if_type inj_is_fun [THEN apply_type])
 apply (simp (no_asm_simp) add: inj_converse_fun [THEN apply_funtype])
@@ -848,7 +848,7 @@
 lemma Un_lepoll_sum: "A \<union> B \<lesssim> A+B"
 apply (unfold lepoll_def)
 apply (rule_tac x = "\<lambda>x\<in>A \<union> B. if x\<in>A then Inl (x) else Inr (x)" in exI)
-apply (rule_tac d = "%z. snd (z)" in lam_injective)
+apply (rule_tac d = "\<lambda>z. snd (z)" in lam_injective)
 apply force
 apply (simp add: Inl_def Inr_def)
 done
@@ -862,7 +862,7 @@
 lemma disj_Un_eqpoll_sum: "A \<inter> B = 0 \<Longrightarrow> A \<union> B \<approx> A + B"
 apply (unfold eqpoll_def)
 apply (rule_tac x = "\<lambda>a\<in>A \<union> B. if a \<in> A then Inl (a) else Inr (a)" in exI)
-apply (rule_tac d = "%z. case (%x. x, %x. x, z)" in lam_bijective)
+apply (rule_tac d = "\<lambda>z. case (\<lambda>x. x, \<lambda>x. x, z)" in lam_bijective)
 apply auto
 done
 
--- a/src/ZF/CardinalArith.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/CardinalArith.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,33 +8,33 @@
 theory CardinalArith imports Cardinal OrderArith ArithSimp Finite begin
 
 definition
-  InfCard       :: "i=>o"  where
+  InfCard       :: "i\<Rightarrow>o"  where
     "InfCard(i) \<equiv> Card(i) \<and> nat \<le> i"
 
 definition
-  cmult         :: "[i,i]=>i"       (infixl \<open>\<otimes>\<close> 70)  where
+  cmult         :: "[i,i]\<Rightarrow>i"       (infixl \<open>\<otimes>\<close> 70)  where
     "i \<otimes> j \<equiv> |i*j|"
 
 definition
-  cadd          :: "[i,i]=>i"       (infixl \<open>\<oplus>\<close> 65)  where
+  cadd          :: "[i,i]\<Rightarrow>i"       (infixl \<open>\<oplus>\<close> 65)  where
     "i \<oplus> j \<equiv> |i+j|"
 
 definition
-  csquare_rel   :: "i=>i"  where
+  csquare_rel   :: "i\<Rightarrow>i"  where
     "csquare_rel(K) \<equiv>
           rvimage(K*K,
-                  lam <x,y>:K*K. <x \<union> y, x, y>,
+                  lam \<langle>x,y\<rangle>:K*K. <x \<union> y, x, y>,
                   rmult(K,Memrel(K), K*K, rmult(K,Memrel(K), K,Memrel(K))))"
 
 definition
-  jump_cardinal :: "i=>i"  where
+  jump_cardinal :: "i\<Rightarrow>i"  where
     \<comment> \<open>This definition is more complex than Kunen's but it more easily proved to
         be a cardinal\<close>
     "jump_cardinal(K) \<equiv>
          \<Union>X\<in>Pow(K). {z. r \<in> Pow(K*K), well_ord(X,r) \<and> z = ordertype(X,r)}"
 
 definition
-  csucc         :: "i=>i"  where
+  csucc         :: "i\<Rightarrow>i"  where
     \<comment> \<open>needed because \<^term>\<open>jump_cardinal(K)\<close> might not be the successor
         of \<^term>\<open>K\<close>\<close>
     "csucc(K) \<equiv> \<mu> L. Card(L) \<and> K<L"
@@ -163,8 +163,8 @@
      "\<lbrakk>A \<lesssim> C;  B \<lesssim> D\<rbrakk> \<Longrightarrow> A + B \<lesssim> C + D"
 apply (unfold lepoll_def)
 apply (elim exE)
-apply (rule_tac x = "\<lambda>z\<in>A+B. case (%w. Inl(f`w), %y. Inr(fa`y), z)" in exI)
-apply (rule_tac d = "case (%w. Inl(converse(f) `w), %y. Inr(converse(fa) ` y))"
+apply (rule_tac x = "\<lambda>z\<in>A+B. case (\<lambda>w. Inl(f`w), \<lambda>y. Inr(fa`y), z)" in exI)
+apply (rule_tac d = "case (\<lambda>w. Inl(converse(f) `w), \<lambda>y. Inr(converse(fa) ` y))"
        in lam_injective)
 apply (typecheck add: inj_is_fun, auto)
 done
@@ -183,8 +183,8 @@
 lemma sum_succ_eqpoll: "succ(A)+B \<approx> succ(A+B)"
 apply (unfold eqpoll_def)
 apply (rule exI)
-apply (rule_tac c = "%z. if z=Inl (A) then A+B else z"
-            and d = "%z. if z=A+B then Inl (A) else z" in lam_bijective)
+apply (rule_tac c = "\<lambda>z. if z=Inl (A) then A+B else z"
+            and d = "\<lambda>z. if z=A+B then Inl (A) else z" in lam_bijective)
    apply simp_all
 apply (blast dest: sym [THEN eq_imp_not_mem] elim: mem_irrefl)+
 done
@@ -221,7 +221,7 @@
 lemma prod_commute_eqpoll: "A*B \<approx> B*A"
 apply (unfold eqpoll_def)
 apply (rule exI)
-apply (rule_tac c = "%<x,y>.<y,x>" and d = "%<x,y>.<y,x>" in lam_bijective,
+apply (rule_tac c = "\<lambda>\<langle>x,y\<rangle>.\<langle>y,x\<rangle>" and d = "\<lambda>\<langle>x,y\<rangle>.\<langle>y,x\<rangle>" in lam_bijective,
        auto)
 done
 
@@ -301,7 +301,7 @@
 
 lemma prod_square_lepoll: "A \<lesssim> A*A"
 apply (unfold lepoll_def inj_def)
-apply (rule_tac x = "\<lambda>x\<in>A. <x,x>" in exI, simp)
+apply (rule_tac x = "\<lambda>x\<in>A. \<langle>x,x\<rangle>" in exI, simp)
 done
 
 (*Could probably weaken the premise to well_ord(K,r), or remove using AC*)
@@ -318,7 +318,7 @@
 
 lemma prod_lepoll_self: "b \<in> B \<Longrightarrow> A \<lesssim> A*B"
 apply (unfold lepoll_def inj_def)
-apply (rule_tac x = "\<lambda>x\<in>A. <x,b>" in exI, simp)
+apply (rule_tac x = "\<lambda>x\<in>A. \<langle>x,b\<rangle>" in exI, simp)
 done
 
 (*Could probably weaken the premises to well_ord(K,r), or removing using AC*)
@@ -337,8 +337,8 @@
      "\<lbrakk>A \<lesssim> C;  B \<lesssim> D\<rbrakk> \<Longrightarrow> A * B  \<lesssim>  C * D"
 apply (unfold lepoll_def)
 apply (elim exE)
-apply (rule_tac x = "lam <w,y>:A*B. <f`w, fa`y>" in exI)
-apply (rule_tac d = "%<w,y>. <converse (f) `w, converse (fa) `y>"
+apply (rule_tac x = "lam \<langle>w,y\<rangle>:A*B. <f`w, fa`y>" in exI)
+apply (rule_tac d = "\<lambda>\<langle>w,y\<rangle>. <converse (f) `w, converse (fa) `y>"
        in lam_injective)
 apply (typecheck add: inj_is_fun, auto)
 done
@@ -357,8 +357,8 @@
 lemma prod_succ_eqpoll: "succ(A)*B \<approx> B + A*B"
 apply (unfold eqpoll_def)
 apply (rule exI)
-apply (rule_tac c = "%<x,y>. if x=A then Inl (y) else Inr (<x,y>)"
-            and d = "case (%y. <A,y>, %z. z)" in lam_bijective)
+apply (rule_tac c = "\<lambda>\<langle>x,y\<rangle>. if x=A then Inl (y) else Inr (\<langle>x,y\<rangle>)"
+            and d = "case (\<lambda>y. \<langle>A,y\<rangle>, \<lambda>z. z)" in lam_bijective)
 apply safe
 apply (simp_all add: succI2 if_type mem_imp_not_eq)
 done
@@ -399,7 +399,7 @@
 
 (*This proof is modelled upon one assuming nat<=A, with injection
   \<lambda>z\<in>cons(u,A). if z=u then 0 else if z \<in> nat then succ(z) else z
-  and inverse %y. if y \<in> nat then nat_case(u, %z. z, y) else y.  \
+  and inverse \<lambda>y. if y \<in> nat then nat_case(u, \<lambda>z. z, y) else y.  \
   If f \<in> inj(nat,A) then range(f) behaves like the natural numbers.*)
 lemma nat_cons_lepoll: "nat \<lesssim> A \<Longrightarrow> cons(u,A) \<lesssim> A"
 apply (unfold lepoll_def)
@@ -410,7 +410,7 @@
              else if z \<in> range (f) then f`succ (converse (f) `z) else z"
        in exI)
 apply (rule_tac d =
-          "%y. if y \<in> range(f) then nat_case (u, %z. f`z, converse(f) `y)
+          "\<lambda>y. if y \<in> range(f) then nat_case (u, \<lambda>z. f`z, converse(f) `y)
                               else y"
        in lam_injective)
 apply (fast intro!: if_type apply_type intro: inj_is_fun inj_converse_fun)
@@ -491,7 +491,7 @@
 subsubsection\<open>Characterising initial segments of the well-ordering\<close>
 
 lemma csquareD:
- "\<lbrakk><<x,y>, <z,z>> \<in> csquare_rel(K);  x<K;  y<K;  z<K\<rbrakk> \<Longrightarrow> x \<le> z \<and> y \<le> z"
+ "\<lbrakk><\<langle>x,y\<rangle>, \<langle>z,z\<rangle>> \<in> csquare_rel(K);  x<K;  y<K;  z<K\<rbrakk> \<Longrightarrow> x \<le> z \<and> y \<le> z"
 apply (unfold csquare_rel_def)
 apply (erule rev_mp)
 apply (elim ltE)
@@ -501,14 +501,14 @@
 done
 
 lemma pred_csquare_subset:
-    "z<K \<Longrightarrow> Order.pred(K*K, <z,z>, csquare_rel(K)) \<subseteq> succ(z)*succ(z)"
+    "z<K \<Longrightarrow> Order.pred(K*K, \<langle>z,z\<rangle>, csquare_rel(K)) \<subseteq> succ(z)*succ(z)"
 apply (unfold Order.pred_def)
 apply (safe del: SigmaI dest!: csquareD)
 apply (unfold lt_def, auto)
 done
 
 lemma csquare_ltI:
- "\<lbrakk>x<z;  y<z;  z<K\<rbrakk> \<Longrightarrow>  <<x,y>, <z,z>> \<in> csquare_rel(K)"
+ "\<lbrakk>x<z;  y<z;  z<K\<rbrakk> \<Longrightarrow>  <\<langle>x,y\<rangle>, \<langle>z,z\<rangle>> \<in> csquare_rel(K)"
 apply (unfold csquare_rel_def)
 apply (subgoal_tac "x<K \<and> y<K")
  prefer 2 apply (blast intro: lt_trans)
@@ -518,7 +518,7 @@
 
 (*Part of the traditional proof.  UNUSED since it's harder to prove \<and> apply *)
 lemma csquare_or_eqI:
- "\<lbrakk>x \<le> z;  y \<le> z;  z<K\<rbrakk> \<Longrightarrow> <<x,y>, <z,z>> \<in> csquare_rel(K) | x=z \<and> y=z"
+ "\<lbrakk>x \<le> z;  y \<le> z;  z<K\<rbrakk> \<Longrightarrow> <\<langle>x,y\<rangle>, \<langle>z,z\<rangle>> \<in> csquare_rel(K) | x=z \<and> y=z"
 apply (unfold csquare_rel_def)
 apply (subgoal_tac "x<K \<and> y<K")
  prefer 2 apply (blast intro: lt_trans1)
@@ -533,8 +533,8 @@
 
 lemma ordermap_z_lt:
       "\<lbrakk>Limit(K);  x<K;  y<K;  z=succ(x \<union> y)\<rbrakk> \<Longrightarrow>
-          ordermap(K*K, csquare_rel(K)) ` <x,y> <
-          ordermap(K*K, csquare_rel(K)) ` <z,z>"
+          ordermap(K*K, csquare_rel(K)) ` \<langle>x,y\<rangle> <
+          ordermap(K*K, csquare_rel(K)) ` \<langle>z,z\<rangle>"
 apply (subgoal_tac "z<K \<and> well_ord (K*K, csquare_rel (K))")
 prefer 2 apply (blast intro!: Un_least_lt Limit_has_succ
                               Limit_is_Ord [THEN well_ord_csquare], clarify)
--- a/src/ZF/Cardinal_AC.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Cardinal_AC.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -118,7 +118,7 @@
 lemma surj_implies_inj:
   assumes f: "f \<in> surj(X,Y)" shows "\<exists>g. g \<in> inj(Y,X)"
 proof -
-  from f AC_Pi [of Y "%y. f-``{y}"]
+  from f AC_Pi [of Y "\<lambda>y. f-``{y}"]
   obtain z where z: "z \<in> (\<Prod>y\<in>Y. f -`` {y})"  
     by (auto simp add: surj_def) (fast dest: apply_Pair)
   show ?thesis
@@ -161,8 +161,8 @@
     proof (unfold lepoll_def)
       show "\<exists>f. f \<in> inj(\<Union>RepFun(K, X), K \<times> K)"
         apply (rule exI) 
-        apply (rule_tac c = "%z. \<langle>\<mu> i. z \<in> X(i), f ` (\<mu> i. z \<in> X(i)) ` z\<rangle>"
-                    and d = "%\<langle>i,j\<rangle>. converse (f`i) ` j" in lam_injective) 
+        apply (rule_tac c = "\<lambda>z. \<langle>\<mu> i. z \<in> X(i), f ` (\<mu> i. z \<in> X(i)) ` z\<rangle>"
+                    and d = "\<lambda>\<langle>i,j\<rangle>. converse (f`i) ` j" in lam_injective) 
         apply (force intro: f inj_is_fun mems apply_type Perm.left_inverse)+
         done
     qed
--- a/src/ZF/Coind/ECR.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Coind/ECR.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -26,7 +26,7 @@
 (* Pointwise extension to environments *)
  
 definition
-  hastyenv :: "[i,i] => o"  where
+  hastyenv :: "[i,i] \<Rightarrow> o"  where
     "hastyenv(ve,te) \<equiv>                        
          ve_dom(ve) = te_dom(te) \<and>          
          (\<forall>x \<in> ve_dom(ve). <ve_app(ve,x),te_app(te,x)> \<in> HasTyRel)"
@@ -55,7 +55,7 @@
        HasTyRel.dom_subset [THEN subsetD, THEN SigmaD1, THEN ValNEE]
 
 lemma hastyenv_owr: 
-  "\<lbrakk>ve \<in> ValEnv; te \<in> TyEnv; hastyenv(ve,te); <v,t> \<in> HasTyRel\<rbrakk> 
+  "\<lbrakk>ve \<in> ValEnv; te \<in> TyEnv; hastyenv(ve,te); \<langle>v,t\<rangle> \<in> HasTyRel\<rbrakk> 
    \<Longrightarrow> hastyenv(ve_owr(ve,x,v),te_owr(te,x,t))"
 by (auto simp add: hastyenv_def ve_app_owr HasTyRel_non_zero)
 
@@ -97,7 +97,7 @@
   "\<lbrakk>ve \<in> ValEnv; x \<in> ExVar; e \<in> Exp; f \<in> ExVar; cl \<in> Val;                
       v_clos(x,e,ve_owr(ve,f,cl)) = cl;                           
       hastyenv(ve,te); <te,e_fix(f,x,e),t> \<in> ElabRel\<rbrakk> \<Longrightarrow>        
-   <cl,t> \<in> HasTyRel"
+   \<langle>cl,t\<rangle> \<in> HasTyRel"
 apply (unfold hastyenv_def)
 apply (erule elab_fixE, safe)
 apply hypsubst_thin
@@ -130,12 +130,12 @@
      \<forall>t te. hastyenv(ve,te) \<longrightarrow>                     
             <te,e1,t> \<in> ElabRel \<longrightarrow> <v_clos(xm,em,vem),t> \<in> HasTyRel;         
      <ve,e2,v2> \<in> EvalRel;                       
-     \<forall>t te. hastyenv(ve,te) \<longrightarrow> <te,e2,t> \<in> ElabRel \<longrightarrow> <v2,t> \<in> HasTyRel;
+     \<forall>t te. hastyenv(ve,te) \<longrightarrow> <te,e2,t> \<in> ElabRel \<longrightarrow> \<langle>v2,t\<rangle> \<in> HasTyRel;
      <ve_owr(vem,xm,v2),em,v> \<in> EvalRel;         
      \<forall>t te. hastyenv(ve_owr(vem,xm,v2),te) \<longrightarrow>      
-            <te,em,t> \<in> ElabRel \<longrightarrow> <v,t> \<in> HasTyRel;
+            <te,em,t> \<in> ElabRel \<longrightarrow> \<langle>v,t\<rangle> \<in> HasTyRel;
      hastyenv(ve,te); <te,e_app(e1,e2),t> \<in> ElabRel\<rbrakk> 
-   \<Longrightarrow> <v,t> \<in> HasTyRel"
+   \<Longrightarrow> \<langle>v,t\<rangle> \<in> HasTyRel"
 apply (erule elab_appE)
 apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+)
 apply (drule spec [THEN spec, THEN mp, THEN mp], assumption+)
@@ -151,7 +151,7 @@
 
 lemma consistency [rule_format]:
    "<ve,e,v> \<in> EvalRel 
-    \<Longrightarrow> (\<forall>t te. hastyenv(ve,te) \<longrightarrow> <te,e,t> \<in> ElabRel \<longrightarrow> <v,t> \<in> HasTyRel)"
+    \<Longrightarrow> (\<forall>t te. hastyenv(ve,te) \<longrightarrow> <te,e,t> \<in> ElabRel \<longrightarrow> \<langle>v,t\<rangle> \<in> HasTyRel)"
 apply (erule EvalRel.induct)
 apply (simp_all add: consistency_const consistency_var consistency_fn 
                      consistency_fix consistency_app1)
--- a/src/ZF/Coind/Language.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Coind/Language.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,7 +10,7 @@
 
 axiomatization
   Const :: i  and               (* Abstract type of constants *)
-  c_app :: "[i,i] => i"         (* Abstract constructor for fun application*)
+  c_app :: "[i,i] \<Rightarrow> i"         (* Abstract constructor for fun application*)
 where
   constNEE:  "c \<in> Const \<Longrightarrow> c \<noteq> 0" and
   c_appI:    "\<lbrakk>c1 \<in> Const; c2 \<in> Const\<rbrakk> \<Longrightarrow> c_app(c1,c2) \<in> Const"
--- a/src/ZF/Coind/Map.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Coind/Map.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,11 +8,11 @@
 theory Map imports ZF begin
 
 definition
-  TMap :: "[i,i] => i"  where
+  TMap :: "[i,i] \<Rightarrow> i"  where
    "TMap(A,B) \<equiv> {m \<in> Pow(A*\<Union>(B)).\<forall>a \<in> A. m``{a} \<in> B}"
 
 definition
-  PMap :: "[i,i] => i"  where
+  PMap :: "[i,i] \<Rightarrow> i"  where
    "PMap(A,B) \<equiv> TMap(A,cons(0,B))"
 
 (* Note: 0 \<in> B \<Longrightarrow> TMap(A,B) = PMap(A,B) *)
@@ -22,11 +22,11 @@
    "map_emp \<equiv> 0"
 
 definition
-  map_owr :: "[i,i,i]=>i"  where
+  map_owr :: "[i,i,i]\<Rightarrow>i"  where
    "map_owr(m,a,b) \<equiv> \<Sum>x \<in> {a} \<union> domain(m). if x=a then b else m``{x}"
 
 definition
-  map_app :: "[i,i]=>i"  where
+  map_app :: "[i,i]\<Rightarrow>i"  where
    "map_app(m,a) \<equiv> m``{a}"
   
 lemma "{0,1} \<subseteq> {1} \<union> TMap(I, {0,1})"
--- a/src/ZF/Coind/Static.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Coind/Static.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,13 +9,13 @@
      parameter of the proof.  A concrete version could be defined inductively.
 ***)
 
-axiomatization isof :: "[i,i] => o"
+axiomatization isof :: "[i,i] \<Rightarrow> o"
   where isof_app: "\<lbrakk>isof(c1,t_fun(t1,t2)); isof(c2,t1)\<rbrakk> \<Longrightarrow> isof(c_app(c1,c2),t2)"
 
 (*Its extension to environments*)
 
 definition
-  isofenv :: "[i,i] => o"  where
+  isofenv :: "[i,i] \<Rightarrow> o"  where
    "isofenv(ve,te) \<equiv>                
       ve_dom(ve) = te_dom(te) \<and>            
       (\<forall>x \<in> ve_dom(ve).                          
--- a/src/ZF/Coind/Types.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Coind/Types.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -24,8 +24,8 @@
           | te_owr ("te \<in> TyEnv","x \<in> ExVar","t \<in> Ty") 
 
 consts
-  te_dom :: "i => i"
-  te_app :: "[i,i] => i"
+  te_dom :: "i \<Rightarrow> i"
+  te_app :: "[i,i] \<Rightarrow> i"
 
 
 primrec (*domain of the type environment*)
--- a/src/ZF/Coind/Values.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Coind/Values.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -21,9 +21,9 @@
   type_intros  A_into_univ mapQU
 
 consts
-  ve_owr :: "[i,i,i] => i"
-  ve_dom :: "i=>i"
-  ve_app :: "[i,i] => i"
+  ve_owr :: "[i,i,i] \<Rightarrow> i"
+  ve_dom :: "i\<Rightarrow>i"
+  ve_app :: "[i,i] \<Rightarrow> i"
 
 
 primrec "ve_owr(ve_mk(m), x, v) = ve_mk(map_owr(m,x,v))"
--- a/src/ZF/Constructible/AC_in_L.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/AC_in_L.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -11,7 +11,7 @@
 text\<open>This could be moved into a library.\<close>
 
 consts
-  rlist   :: "[i,i]=>i"
+  rlist   :: "[i,i]\<Rightarrow>i"
 
 inductive
   domains "rlist(A,r)" \<subseteq> "list(A) * list(A)"
@@ -159,20 +159,20 @@
   assumes fn_inj: "fn \<in> inj(nat*nat, nat)"
 
 
-consts   enum :: "[i,i]=>i"
+consts   enum :: "[i,i]\<Rightarrow>i"
 primrec
-  "enum(f, Member(x,y)) = f ` <0, f ` <x,y>>"
-  "enum(f, Equal(x,y)) = f ` <1, f ` <x,y>>"
+  "enum(f, Member(x,y)) = f ` <0, f ` \<langle>x,y\<rangle>>"
+  "enum(f, Equal(x,y)) = f ` <1, f ` \<langle>x,y\<rangle>>"
   "enum(f, Nand(p,q)) = f ` <2, f ` <enum(f,p), enum(f,q)>>"
   "enum(f, Forall(p)) = f ` <succ(2), enum(f,p)>"
 
 lemma (in Nat_Times_Nat) fn_type [TC,simp]:
-    "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> fn`<x,y> \<in> nat"
+    "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> fn`\<langle>x,y\<rangle> \<in> nat"
 by (blast intro: inj_is_fun [OF fn_inj] apply_funtype)
 
 lemma (in Nat_Times_Nat) fn_iff:
     "\<lbrakk>x \<in> nat; y \<in> nat; u \<in> nat; v \<in> nat\<rbrakk>
-     \<Longrightarrow> (fn`<x,y> = fn`<u,v>) \<longleftrightarrow> (x=u \<and> y=v)"
+     \<Longrightarrow> (fn`\<langle>x,y\<rangle> = fn`\<langle>u,v\<rangle>) \<longleftrightarrow> (x=u \<and> y=v)"
 by (blast dest: inj_apply_equality [OF fn_inj])
 
 lemma (in Nat_Times_Nat) enum_type [TC,simp]:
@@ -230,34 +230,34 @@
 (environment, formula) pairs into the ordinals.  For each member of \<^term>\<open>DPow(A)\<close>, we take the minimum such ordinal.\<close>
 
 definition
-  env_form_r :: "[i,i,i]=>i" where
+  env_form_r :: "[i,i,i]\<Rightarrow>i" where
     \<comment> \<open>wellordering on (environment, formula) pairs\<close>
    "env_form_r(f,r,A) \<equiv>
       rmult(list(A), rlist(A, r),
             formula, measure(formula, enum(f)))"
 
 definition
-  env_form_map :: "[i,i,i,i]=>i" where
+  env_form_map :: "[i,i,i,i]\<Rightarrow>i" where
     \<comment> \<open>map from (environment, formula) pairs to ordinals\<close>
    "env_form_map(f,r,A,z)
       \<equiv> ordermap(list(A) * formula, env_form_r(f,r,A)) ` z"
 
 definition
-  DPow_ord :: "[i,i,i,i,i]=>o" where
+  DPow_ord :: "[i,i,i,i,i]\<Rightarrow>o" where
     \<comment> \<open>predicate that holds if \<^term>\<open>k\<close> is a valid index for \<^term>\<open>X\<close>\<close>
    "DPow_ord(f,r,A,X,k) \<equiv>
            \<exists>env \<in> list(A). \<exists>p \<in> formula.
              arity(p) \<le> succ(length(env)) \<and>
              X = {x\<in>A. sats(A, p, Cons(x,env))} \<and>
-             env_form_map(f,r,A,<env,p>) = k"
+             env_form_map(f,r,A,\<langle>env,p\<rangle>) = k"
 
 definition
-  DPow_least :: "[i,i,i,i]=>i" where
+  DPow_least :: "[i,i,i,i]\<Rightarrow>i" where
     \<comment> \<open>function yielding the smallest index for \<^term>\<open>X\<close>\<close>
    "DPow_least(f,r,A,X) \<equiv> \<mu> k. DPow_ord(f,r,A,X,k)"
 
 definition
-  DPow_r :: "[i,i,i]=>i" where
+  DPow_r :: "[i,i,i]\<Rightarrow>i" where
     \<comment> \<open>a wellordering on \<^term>\<open>DPow(A)\<close>\<close>
    "DPow_r(f,r,A) \<equiv> measure(DPow(A), DPow_least(f,r,A))"
 
@@ -329,7 +329,7 @@
 of wellorderings for smaller ordinals.\<close>
 
 definition
-  rlimit :: "[i,i=>i]=>i" where
+  rlimit :: "[i,i\<Rightarrow>i]\<Rightarrow>i" where
   \<comment> \<open>Expresses the wellordering at limit ordinals.  The conditional
       lets us remove the premise \<^term>\<open>Limit(i)\<close> from some theorems.\<close>
     "rlimit(i,r) \<equiv>
@@ -341,7 +341,7 @@
        else 0"
 
 definition
-  Lset_new :: "i=>i" where
+  Lset_new :: "i\<Rightarrow>i" where
   \<comment> \<open>This constant denotes the set of elements introduced at level
       \<^term>\<open>succ(i)\<close>\<close>
     "Lset_new(i) \<equiv> {x \<in> Lset(succ(i)). lrank(x) = i}"
@@ -413,8 +413,8 @@
 subsection\<open>Transfinite Definition of the Wellordering on \<^term>\<open>L\<close>\<close>
 
 definition
-  L_r :: "[i, i] => i" where
-  "L_r(f) \<equiv> %i.
+  L_r :: "[i, i] \<Rightarrow> i" where
+  "L_r(f) \<equiv> \<lambda>i.
       transrec3(i, 0, \<lambda>x r. DPow_r(f, r, Lset(x)), 
                 \<lambda>x r. rlimit(x, \<lambda>y. r`y))"
 
--- a/src/ZF/Constructible/DPow_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/DPow_absolute.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -15,7 +15,7 @@
 text\<open>The three arguments of \<^term>\<open>p\<close> are always 2, 1, 0.  It is buried
    within 11 quantifiers\<And>\<close>
 
-(* is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o"
+(* is_formula_rec :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i] \<Rightarrow> o"
    "is_formula_rec(M,MH,p,z)  \<equiv>
       \<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) \<and> is_depth(M,p,dp) \<and> 
        2       1      0
@@ -23,7 +23,7 @@
 *)
 
 definition
-  formula_rec_fm :: "[i, i, i]=>i" where
+  formula_rec_fm :: "[i, i, i]\<Rightarrow>i" where
  "formula_rec_fm(mh,p,z) \<equiv> 
     Exists(Exists(Exists(
       And(finite_ordinal_fm(2),
@@ -81,7 +81,7 @@
 
 (* is_satisfies(M,A,p,z) \<equiv> is_formula_rec (M, satisfies_MH(M,A), p, z) *)
 definition
-  satisfies_fm :: "[i,i,i]=>i" where
+  satisfies_fm :: "[i,i,i]\<Rightarrow>i" where
     "satisfies_fm(x) \<equiv> formula_rec_fm (satisfies_MH_fm(x#+5#+6, 2, 1, 0))"
 
 lemma is_satisfies_type [TC]:
@@ -113,14 +113,14 @@
 lemma DPow'_eq: 
   "DPow'(A) = {z . ep \<in> list(A) * formula, 
                     \<exists>env \<in> list(A). \<exists>p \<in> formula. 
-                       ep = <env,p> \<and> z = {x\<in>A. sats(A, p, Cons(x,env))}}"
+                       ep = \<langle>env,p\<rangle> \<and> z = {x\<in>A. sats(A, p, Cons(x,env))}}"
 by (simp add: DPow'_def, blast) 
 
 
 text\<open>Relativize the use of \<^term>\<open>sats\<close> within \<^term>\<open>DPow'\<close>
 (the comprehension).\<close>
 definition
-  is_DPow_sats :: "[i=>o,i,i,i,i] => o" where
+  is_DPow_sats :: "[i\<Rightarrow>o,i,i,i,i] \<Rightarrow> o" where
    "is_DPow_sats(M,A,env,p,x) \<equiv>
       \<forall>n1[M]. \<forall>e[M]. \<forall>sp[M]. 
              is_satisfies(M,A,p,sp) \<longrightarrow> is_Cons(M,x,env,e) \<longrightarrow> 
@@ -149,7 +149,7 @@
              fun_apply(M, sp, e, n1) \<longrightarrow> number1(M, n1) *)
 
 definition
-  DPow_sats_fm :: "[i,i,i,i]=>i" where
+  DPow_sats_fm :: "[i,i,i,i]\<Rightarrow>i" where
   "DPow_sats_fm(A,env,p,x) \<equiv>
    Forall(Forall(Forall(
      Implies(satisfies_fm(A#+3,p#+3,0), 
@@ -205,7 +205,7 @@
    "M(A)
     \<Longrightarrow> strong_replacement (M, 
          \<lambda>ep z. \<exists>env\<in>list(A). \<exists>p\<in>formula.
-                  ep = <env,p> \<and> z = {x \<in> A . sats(A, p, Cons(x, env))})" 
+                  ep = \<langle>env,p\<rangle> \<and> z = {x \<in> A . sats(A, p, Cons(x, env))})" 
 by (insert rep [of A], simp add: Collect_DPow_sats_abs) 
 
 
@@ -220,7 +220,7 @@
 
 text\<open>Relativization of the Operator \<^term>\<open>DPow'\<close>\<close>
 definition 
-  is_DPow' :: "[i=>o,i,i] => o" where
+  is_DPow' :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_DPow'(M,A,Z) \<equiv> 
        \<forall>X[M]. X \<in> Z \<longleftrightarrow> 
          subset(M,X,A) \<and> 
@@ -308,11 +308,11 @@
 text\<open>The formula \<^term>\<open>is_P\<close> has one free variable, 0, and it is
 enclosed within a single quantifier.\<close>
 
-(* is_Collect :: "[i=>o,i,i=>o,i] => o"
+(* is_Collect :: "[i\<Rightarrow>o,i,i\<Rightarrow>o,i] \<Rightarrow> o"
     "is_Collect(M,A,P,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A \<and> P(x)" *)
 
 definition
-  Collect_fm :: "[i, i, i]=>i" where
+  Collect_fm :: "[i, i, i]\<Rightarrow>i" where
  "Collect_fm(A,is_P,z) \<equiv> 
         Forall(Iff(Member(0,succ(z)),
                    And(Member(0,succ(A)), is_P)))"
@@ -359,11 +359,11 @@
 text\<open>BEWARE!  The formula \<^term>\<open>is_P\<close> has free variables 0, 1
  and not the usual 1, 0!  It is enclosed within two quantifiers.\<close>
 
-(*  is_Replace :: "[i=>o,i,[i,i]=>o,i] => o"
+(*  is_Replace :: "[i\<Rightarrow>o,i,[i,i]\<Rightarrow>o,i] \<Rightarrow> o"
     "is_Replace(M,A,P,z) \<equiv> \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A \<and> P(x,u))" *)
 
 definition
-  Replace_fm :: "[i, i, i]=>i" where
+  Replace_fm :: "[i, i, i]\<Rightarrow>i" where
   "Replace_fm(A,is_P,z) \<equiv> 
         Forall(Iff(Member(0,succ(z)),
                    Exists(And(Member(0,A#+2), is_P))))"
@@ -417,7 +417,7 @@
                     is_Collect(M, A, is_DPow_sats(M,A,env,p), X))" *)
 
 definition
-  DPow'_fm :: "[i,i]=>i" where
+  DPow'_fm :: "[i,i]\<Rightarrow>i" where
     "DPow'_fm(A,Z) \<equiv> 
       Forall(
        Iff(Member(0,succ(Z)),
@@ -456,7 +456,7 @@
 subsection\<open>A Locale for Relativizing the Operator \<^term>\<open>Lset\<close>\<close>
 
 definition
-  transrec_body :: "[i=>o,i,i,i,i] => o" where
+  transrec_body :: "[i\<Rightarrow>o,i,i,i,i] \<Rightarrow> o" where
     "transrec_body(M,g,x) \<equiv>
       \<lambda>y z. \<exists>gy[M]. y \<in> x \<and> fun_apply(M,g,y,gy) \<and> is_DPow'(M,gy,z)"
 
@@ -508,11 +508,11 @@
 text\<open>Relativization of the Operator \<^term>\<open>Lset\<close>\<close>
 
 definition
-  is_Lset :: "[i=>o, i, i] => o" where
+  is_Lset :: "[i\<Rightarrow>o, i, i] \<Rightarrow> o" where
    \<comment> \<open>We can use the term language below because \<^term>\<open>is_Lset\<close> will
        not have to be internalized: it isn't used in any instance of
        separation.\<close>
-   "is_Lset(M,a,z) \<equiv> is_transrec(M, %x f u. u = (\<Union>y\<in>x. DPow'(f`y)), a, z)"
+   "is_Lset(M,a,z) \<equiv> is_transrec(M, \<lambda>x f u. u = (\<Union>y\<in>x. DPow'(f`y)), a, z)"
 
 lemma (in M_Lset) Lset_abs:
   "\<lbrakk>Ord(i);  M(i);  M(z)\<rbrakk> 
@@ -613,7 +613,7 @@
 subsection\<open>The Notion of Constructible Set\<close>
 
 definition
-  constructible :: "[i=>o,i] => o" where
+  constructible :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "constructible(M,x) \<equiv>
        \<exists>i[M]. \<exists>Li[M]. ordinal(M,i) \<and> is_Lset(M,i,Li) \<and> x \<in> Li"
 
--- a/src/ZF/Constructible/Datatype_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Datatype_absolute.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,11 +10,11 @@
 subsection\<open>The lfp of a continuous function can be expressed as a union\<close>
 
 definition
-  directed :: "i=>o" where
+  directed :: "i\<Rightarrow>o" where
    "directed(A) \<equiv> A\<noteq>0 \<and> (\<forall>x\<in>A. \<forall>y\<in>A. x \<union> y \<in> A)"
 
 definition
-  contin :: "(i=>i) => o" where
+  contin :: "(i\<Rightarrow>i) \<Rightarrow> o" where
    "contin(h) \<equiv> (\<forall>A. directed(A) \<longrightarrow> h(\<Union>A) = (\<Union>X\<in>A. h(X)))"
 
 lemma bnd_mono_iterates_subset: "\<lbrakk>bnd_mono(D, h); n \<in> nat\<rbrakk> \<Longrightarrow> h^n (0) \<subseteq> D"
@@ -114,19 +114,19 @@
 subsection \<open>Absoluteness for "Iterates"\<close>
 
 definition
-  iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where
+  iterates_MH :: "[i\<Rightarrow>o, [i,i]\<Rightarrow>o, i, i, i, i] \<Rightarrow> o" where
    "iterates_MH(M,isF,v,n,g,z) \<equiv>
         is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) \<and> isF(gm,u),
                     n, z)"
 
 definition
-  is_iterates :: "[i=>o, [i,i]=>o, i, i, i] => o" where
+  is_iterates :: "[i\<Rightarrow>o, [i,i]\<Rightarrow>o, i, i, i] \<Rightarrow> o" where
     "is_iterates(M,isF,v,n,Z) \<equiv> 
       \<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) \<and> membership(M,sn,msn) \<and>
                        is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"
 
 definition
-  iterates_replacement :: "[i=>o, [i,i]=>o, i] => o" where
+  iterates_replacement :: "[i\<Rightarrow>o, [i,i]\<Rightarrow>o, i] \<Rightarrow> o" where
    "iterates_replacement(M,isF,v) \<equiv>
       \<forall>n[M]. n\<in>nat \<longrightarrow> 
          wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))"
@@ -209,7 +209,7 @@
 
 
 definition
-  is_list_functor :: "[i=>o,i,i,i] => o" where
+  is_list_functor :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_list_functor(M,A,X,Z) \<equiv> 
         \<exists>n1[M]. \<exists>AX[M]. 
          number1(M,n1) \<and> cartprod(M,A,X,AX) \<and> is_sum(M,n1,AX,Z)"
@@ -262,7 +262,7 @@
 
 
 definition
-  is_formula_functor :: "[i=>o,i,i] => o" where
+  is_formula_functor :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_formula_functor(M,X,Z) \<equiv> 
         \<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M]. 
           omega(M,nat') \<and> cartprod(M,nat',nat',natnat) \<and> 
@@ -280,7 +280,7 @@
 subsection\<open>\<^term>\<open>M\<close> Contains the List and Formula Datatypes\<close>
 
 definition
-  list_N :: "[i,i] => i" where
+  list_N :: "[i,i] \<Rightarrow> i" where
     "list_N(A,n) \<equiv> (\<lambda>X. {0} + A * X)^n (0)"
 
 lemma Nil_in_list_N [simp]: "[] \<in> list_N(A,succ(n))"
@@ -341,24 +341,24 @@
 done
 
 definition
-  is_list_N :: "[i=>o,i,i,i] => o" where
+  is_list_N :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_list_N(M,A,n,Z) \<equiv> 
       \<exists>zero[M]. empty(M,zero) \<and> 
                 is_iterates(M, is_list_functor(M,A), zero, n, Z)"
 
 definition  
-  mem_list :: "[i=>o,i,i] => o" where
+  mem_list :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "mem_list(M,A,l) \<equiv> 
       \<exists>n[M]. \<exists>listn[M]. 
        finite_ordinal(M,n) \<and> is_list_N(M,A,n,listn) \<and> l \<in> listn"
 
 definition
-  is_list :: "[i=>o,i,i] => o" where
+  is_list :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_list(M,A,Z) \<equiv> \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l)"
 
 subsubsection\<open>Towards Absoluteness of \<^term>\<open>formula_rec\<close>\<close>
 
-consts   depth :: "i=>i"
+consts   depth :: "i\<Rightarrow>i"
 primrec
   "depth(Member(x,y)) = 0"
   "depth(Equal(x,y))  = 0"
@@ -370,7 +370,7 @@
 
 
 definition
-  formula_N :: "i => i" where
+  formula_N :: "i \<Rightarrow> i" where
     "formula_N(n) \<equiv> (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)"
 
 lemma Member_in_formula_N [simp]:
@@ -445,20 +445,20 @@
 done
 
 definition
-  is_formula_N :: "[i=>o,i,i] => o" where
+  is_formula_N :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_formula_N(M,n,Z) \<equiv> 
       \<exists>zero[M]. empty(M,zero) \<and> 
                 is_iterates(M, is_formula_functor(M), zero, n, Z)"
 
 
 definition  
-  mem_formula :: "[i=>o,i] => o" where
+  mem_formula :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "mem_formula(M,p) \<equiv> 
       \<exists>n[M]. \<exists>formn[M]. 
        finite_ordinal(M,n) \<and> is_formula_N(M,n,formn) \<and> p \<in> formn"
 
 definition
-  is_formula :: "[i=>o,i] => o" where
+  is_formula :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "is_formula(M,Z) \<equiv> \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p)"
 
 locale M_datatypes = M_trancl +
@@ -473,7 +473,7 @@
    "strong_replacement(M,
          \<lambda>n y. n\<in>nat \<and> is_iterates(M, is_formula_functor(M), 0, n, y))"
   and nth_replacement:
-   "M(l) \<Longrightarrow> iterates_replacement(M, %l t. is_tl(M,l,t), l)"
+   "M(l) \<Longrightarrow> iterates_replacement(M, \<lambda>l t. is_tl(M,l,t), l)"
 
 
 subsubsection\<open>Absoluteness of the List Construction\<close>
@@ -588,17 +588,17 @@
 done
 
 definition
-  is_eclose_n :: "[i=>o,i,i,i] => o" where
+  is_eclose_n :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_eclose_n(M,A,n,Z) \<equiv> is_iterates(M, big_union(M), A, n, Z)"
 
 definition
-  mem_eclose :: "[i=>o,i,i] => o" where
+  mem_eclose :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "mem_eclose(M,A,l) \<equiv>
       \<exists>n[M]. \<exists>eclosen[M].
        finite_ordinal(M,n) \<and> is_eclose_n(M,A,n,eclosen) \<and> l \<in> eclosen"
 
 definition
-  is_eclose :: "[i=>o,i,i] => o" where
+  is_eclose :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_eclose(M,A,Z) \<equiv> \<forall>u[M]. u \<in> Z \<longleftrightarrow> mem_eclose(M,A,u)"
 
 
@@ -650,14 +650,14 @@
 text\<open>\<^prop>\<open>transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)\<close>\<close>
 
 definition
-  is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
+  is_transrec :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i] \<Rightarrow> o" where
    "is_transrec(M,MH,a,z) \<equiv>
       \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
        upair(M,a,a,sa) \<and> is_eclose(M,sa,esa) \<and> membership(M,esa,mesa) \<and>
        is_wfrec(M,MH,mesa,a,z)"
 
 definition
-  transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where
+  transrec_replacement :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i] \<Rightarrow> o" where
    "transrec_replacement(M,MH,a) \<equiv>
       \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
        upair(M,a,a,sa) \<and> is_eclose(M,sa,esa) \<and> membership(M,esa,mesa) \<and>
@@ -698,7 +698,7 @@
 text\<open>But it is never used.\<close>
 
 definition
-  is_length :: "[i=>o,i,i,i] => o" where
+  is_length :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_length(M,A,l,n) \<equiv>
        \<exists>sn[M]. \<exists>list_n[M]. \<exists>list_sn[M].
         is_list_N(M,A,n,list_n) \<and> l \<notin> list_n \<and>
@@ -746,7 +746,7 @@
 done
 
 definition
-  is_nth :: "[i=>o,i,i,i] => o" where
+  is_nth :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_nth(M,n,l,Z) \<equiv>
       \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) \<and> is_hd(M,X,Z)"
 
@@ -764,7 +764,7 @@
 subsection\<open>Relativization and Absoluteness for the \<^term>\<open>formula\<close> Constructors\<close>
 
 definition
-  is_Member :: "[i=>o,i,i,i] => o" where
+  is_Member :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
      \<comment> \<open>because \<^term>\<open>Member(x,y) \<equiv> Inl(Inl(\<langle>x,y\<rangle>))\<close>\<close>
     "is_Member(M,x,y,Z) \<equiv>
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inl(M,p,u) \<and> is_Inl(M,u,Z)"
@@ -778,7 +778,7 @@
 by (simp add: Member_def)
 
 definition
-  is_Equal :: "[i=>o,i,i,i] => o" where
+  is_Equal :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
      \<comment> \<open>because \<^term>\<open>Equal(x,y) \<equiv> Inl(Inr(\<langle>x,y\<rangle>))\<close>\<close>
     "is_Equal(M,x,y,Z) \<equiv>
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inr(M,p,u) \<and> is_Inl(M,u,Z)"
@@ -791,7 +791,7 @@
 by (simp add: Equal_def)
 
 definition
-  is_Nand :: "[i=>o,i,i,i] => o" where
+  is_Nand :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
      \<comment> \<open>because \<^term>\<open>Nand(x,y) \<equiv> Inr(Inl(\<langle>x,y\<rangle>))\<close>\<close>
     "is_Nand(M,x,y,Z) \<equiv>
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inl(M,p,u) \<and> is_Inr(M,u,Z)"
@@ -804,7 +804,7 @@
 by (simp add: Nand_def)
 
 definition
-  is_Forall :: "[i=>o,i,i] => o" where
+  is_Forall :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
      \<comment> \<open>because \<^term>\<open>Forall(x) \<equiv> Inr(Inr(p))\<close>\<close>
     "is_Forall(M,p,Z) \<equiv> \<exists>u[M]. is_Inr(M,p,u) \<and> is_Inr(M,u,Z)"
 
@@ -820,7 +820,7 @@
 subsection \<open>Absoluteness for \<^term>\<open>formula_rec\<close>\<close>
 
 definition
-  formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" where
+  formula_rec_case :: "[[i,i]\<Rightarrow>i, [i,i]\<Rightarrow>i, [i,i,i,i]\<Rightarrow>i, [i,i]\<Rightarrow>i, i, i] \<Rightarrow> i" where
     \<comment> \<open>the instance of \<^term>\<open>formula_case\<close> in \<^term>\<open>formula_rec\<close>\<close>
    "formula_rec_case(a,b,c,d,h) \<equiv>
         formula_case (a, b,
@@ -854,7 +854,7 @@
 subsubsection\<open>Absoluteness for the Formula Operator \<^term>\<open>depth\<close>\<close>
 
 definition
-  is_depth :: "[i=>o,i,i] => o" where
+  is_depth :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_depth(M,p,n) \<equiv>
        \<exists>sn[M]. \<exists>formula_n[M]. \<exists>formula_sn[M].
         is_formula_N(M,n,formula_n) \<and> p \<notin> formula_n \<and>
@@ -880,7 +880,7 @@
 
 definition
  is_formula_case ::
-    "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where
+    "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, [i,i,i]\<Rightarrow>o, [i,i,i]\<Rightarrow>o, [i,i]\<Rightarrow>o, i, i] \<Rightarrow> o" where
   \<comment> \<open>no constraint on non-formulas\<close>
   "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) \<equiv>
       (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) \<longrightarrow> finite_ordinal(M,y) \<longrightarrow>
@@ -914,7 +914,7 @@
 subsubsection \<open>Absoluteness for \<^term>\<open>formula_rec\<close>: Final Results\<close>
 
 definition
-  is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
+  is_formula_rec :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i] \<Rightarrow> o" where
     \<comment> \<open>predicate to relativize the functional \<^term>\<open>formula_rec\<close>\<close>
    "is_formula_rec(M,MH,p,z)  \<equiv>
       \<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) \<and> is_depth(M,p,dp) \<and>
--- a/src/ZF/Constructible/Formula.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Formula.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -21,27 +21,27 @@
 declare formula.intros [TC]
 
 definition
-  Neg :: "i=>i" where
+  Neg :: "i\<Rightarrow>i" where
   "Neg(p) \<equiv> Nand(p,p)"
 
 definition
-  And :: "[i,i]=>i" where
+  And :: "[i,i]\<Rightarrow>i" where
   "And(p,q) \<equiv> Neg(Nand(p,q))"
 
 definition
-  Or :: "[i,i]=>i" where
+  Or :: "[i,i]\<Rightarrow>i" where
   "Or(p,q) \<equiv> Nand(Neg(p),Neg(q))"
 
 definition
-  Implies :: "[i,i]=>i" where
+  Implies :: "[i,i]\<Rightarrow>i" where
   "Implies(p,q) \<equiv> Nand(p,Neg(q))"
 
 definition
-  Iff :: "[i,i]=>i" where
+  Iff :: "[i,i]\<Rightarrow>i" where
   "Iff(p,q) \<equiv> And(Implies(p,q), Implies(q,p))"
 
 definition
-  Exists :: "i=>i" where
+  Exists :: "i\<Rightarrow>i" where
   "Exists(p) \<equiv> Neg(Forall(Neg(p)))"
 
 lemma Neg_type [TC]: "p \<in> formula \<Longrightarrow> Neg(p) \<in> formula"
@@ -65,7 +65,7 @@
 by (simp add: Exists_def)
 
 
-consts   satisfies :: "[i,i]=>i"
+consts   satisfies :: "[i,i]\<Rightarrow>i"
 primrec (*explicit lambda is required because the environment varies*)
   "satisfies(A,Member(x,y)) =
       (\<lambda>env \<in> list(A). bool_of_o (nth(x,env) \<in> nth(y,env)))"
@@ -84,7 +84,7 @@
 by (induct set: formula) simp_all
 
 abbreviation
-  sats :: "[i,i,i] => o" where
+  sats :: "[i,i,i] \<Rightarrow> o" where
   "sats(A,p,env) \<equiv> satisfies(A,p)`env = 1"
 
 lemma sats_Member_iff [simp]:
@@ -195,7 +195,7 @@
 
 subsection\<open>Arity of a Formula: Maximum Free de Bruijn Index\<close>
 
-consts   arity :: "i=>i"
+consts   arity :: "i\<Rightarrow>i"
 primrec
   "arity(Member(x,y)) = succ(x) \<union> succ(y)"
 
@@ -250,7 +250,7 @@
 subsection\<open>Renaming Some de Bruijn Variables\<close>
 
 definition
-  incr_var :: "[i,i]=>i" where
+  incr_var :: "[i,i]\<Rightarrow>i" where
   "incr_var(x,nq) \<equiv> if x<nq then x else succ(x)"
 
 lemma incr_var_lt: "x<nq \<Longrightarrow> incr_var(x,nq) = x"
@@ -261,7 +261,7 @@
 apply (blast dest: lt_trans1)
 done
 
-consts   incr_bv :: "i=>i"
+consts   incr_bv :: "i\<Rightarrow>i"
 primrec
   "incr_bv(Member(x,y)) =
       (\<lambda>nq \<in> nat. Member (incr_var(x,nq), incr_var(y,nq)))"
@@ -339,7 +339,7 @@
 subsection\<open>Renaming all but the First de Bruijn Variable\<close>
 
 definition
-  incr_bv1 :: "i => i" where
+  incr_bv1 :: "i \<Rightarrow> i" where
   "incr_bv1(p) \<equiv> incr_bv(p)`1"
 
 
@@ -391,7 +391,7 @@
 
 text\<open>The definable powerset operation: Kunen's definition VI 1.1, page 165.\<close>
 definition
-  DPow :: "i => i" where
+  DPow :: "i \<Rightarrow> i" where
   "DPow(A) \<equiv> {X \<in> Pow(A).
                \<exists>env \<in> list(A). \<exists>p \<in> formula.
                  arity(p) \<le> succ(length(env)) \<and>
@@ -514,7 +514,7 @@
 subsubsection\<open>The subset relation\<close>
 
 definition
-  subset_fm :: "[i,i]=>i" where
+  subset_fm :: "[i,i]\<Rightarrow>i" where
   "subset_fm(x,y) \<equiv> Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
 
 lemma subset_type [TC]: "\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> subset_fm(x,y) \<in> formula"
@@ -535,7 +535,7 @@
 subsubsection\<open>Transitive sets\<close>
 
 definition
-  transset_fm :: "i=>i" where
+  transset_fm :: "i\<Rightarrow>i" where
   "transset_fm(x) \<equiv> Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
 
 lemma transset_type [TC]: "x \<in> nat \<Longrightarrow> transset_fm(x) \<in> formula"
@@ -556,7 +556,7 @@
 subsubsection\<open>Ordinals\<close>
 
 definition
-  ordinal_fm :: "i=>i" where
+  ordinal_fm :: "i\<Rightarrow>i" where
   "ordinal_fm(x) \<equiv>
     And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))"
 
@@ -588,11 +588,11 @@
 subsection\<open>Constant Lset: Levels of the Constructible Universe\<close>
 
 definition
-  Lset :: "i=>i" where
-  "Lset(i) \<equiv> transrec(i, %x f. \<Union>y\<in>x. DPow(f`y))"
+  Lset :: "i\<Rightarrow>i" where
+  "Lset(i) \<equiv> transrec(i, \<lambda>x f. \<Union>y\<in>x. DPow(f`y))"
 
 definition
-  L :: "i=>o" where \<comment> \<open>Kunen's definition VI 1.5, page 167\<close>
+  L :: "i\<Rightarrow>o" where \<comment> \<open>Kunen's definition VI 1.5, page 167\<close>
   "L(x) \<equiv> \<exists>i. Ord(i) \<and> x \<in> Lset(i)"
 
 text\<open>NOT SUITABLE FOR REWRITING -- RECURSIVE!\<close>
@@ -798,7 +798,7 @@
 by (simp add: Lset_succ empty_in_DPow cons_in_DPow)
 
 lemma Pair_in_Lset:
-    "\<lbrakk>a \<in> Lset(i);  b \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Lset(succ(succ(i)))"
+    "\<lbrakk>a \<in> Lset(i);  b \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Lset(succ(succ(i)))"
 apply (unfold Pair_def)
 apply (blast intro: doubleton_in_Lset)
 done
@@ -822,7 +822,7 @@
 done
 
 lemma Pair_in_LLimit:
-    "\<lbrakk>a \<in> Lset(i);  b \<in> Lset(i);  Limit(i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Lset(i)"
+    "\<lbrakk>a \<in> Lset(i);  b \<in> Lset(i);  Limit(i)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Lset(i)"
 txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close>
 apply (erule Limit_LsetE, assumption)
 apply (erule Limit_LsetE, assumption)
@@ -835,7 +835,7 @@
 
 text\<open>The rank function for the constructible universe\<close>
 definition
-  lrank :: "i=>i" where \<comment> \<open>Kunen's definition VI 1.7\<close>
+  lrank :: "i\<Rightarrow>i" where \<comment> \<open>Kunen's definition VI 1.7\<close>
   "lrank(x) \<equiv> \<mu> i. x \<in> Lset(succ(i))"
 
 lemma L_I: "\<lbrakk>x \<in> Lset(i); Ord(i)\<rbrakk> \<Longrightarrow> L(x)"
@@ -995,7 +995,7 @@
 
 text\<open>A simpler version of \<^term>\<open>DPow\<close>: no arity check!\<close>
 definition
-  DPow' :: "i => i" where
+  DPow' :: "i \<Rightarrow> i" where
   "DPow'(A) \<equiv> {X \<in> Pow(A).
                 \<exists>env \<in> list(A). \<exists>p \<in> formula.
                     X = {x\<in>A. sats(A, p, Cons(x,env))}}"
@@ -1022,7 +1022,7 @@
 
 text\<open>And thus we can relativize \<^term>\<open>Lset\<close> without bothering with
       \<^term>\<open>arity\<close> and \<^term>\<open>length\<close>\<close>
-lemma Lset_eq_transrec_DPow': "Lset(i) = transrec(i, %x f. \<Union>y\<in>x. DPow'(f`y))"
+lemma Lset_eq_transrec_DPow': "Lset(i) = transrec(i, \<lambda>x f. \<Union>y\<in>x. DPow'(f`y))"
 apply (rule_tac a=i in eps_induct)
 apply (subst Lset)
 apply (subst transrec)
--- a/src/ZF/Constructible/Internalize.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Internalize.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,7 +10,7 @@
 
 (*  is_Inl(M,a,z) \<equiv> \<exists>zero[M]. empty(M,zero) \<and> pair(M,zero,a,z) *)
 definition
-  Inl_fm :: "[i,i]=>i" where
+  Inl_fm :: "[i,i]\<Rightarrow>i" where
     "Inl_fm(a,z) \<equiv> Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))"
 
 lemma Inl_type [TC]:
@@ -40,7 +40,7 @@
 
 (*  is_Inr(M,a,z) \<equiv> \<exists>n1[M]. number1(M,n1) \<and> pair(M,n1,a,z) *)
 definition
-  Inr_fm :: "[i,i]=>i" where
+  Inr_fm :: "[i,i]\<Rightarrow>i" where
     "Inr_fm(a,z) \<equiv> Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))"
 
 lemma Inr_type [TC]:
@@ -71,7 +71,7 @@
 (* is_Nil(M,xs) \<equiv> \<exists>zero[M]. empty(M,zero) \<and> is_Inl(M,zero,xs) *)
 
 definition
-  Nil_fm :: "i=>i" where
+  Nil_fm :: "i\<Rightarrow>i" where
     "Nil_fm(x) \<equiv> Exists(And(empty_fm(0), Inl_fm(0,succ(x))))"
 
 lemma Nil_type [TC]: "x \<in> nat \<Longrightarrow> Nil_fm(x) \<in> formula"
@@ -100,7 +100,7 @@
 
 (*  "is_Cons(M,a,l,Z) \<equiv> \<exists>p[M]. pair(M,a,l,p) \<and> is_Inr(M,p,Z)" *)
 definition
-  Cons_fm :: "[i,i,i]=>i" where
+  Cons_fm :: "[i,i,i]\<Rightarrow>i" where
     "Cons_fm(a,l,Z) \<equiv>
        Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))"
 
@@ -132,7 +132,7 @@
 (* is_quasilist(M,xs) \<equiv> is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))" *)
 
 definition
-  quasilist_fm :: "i=>i" where
+  quasilist_fm :: "i\<Rightarrow>i" where
     "quasilist_fm(x) \<equiv>
        Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))"
 
@@ -167,7 +167,7 @@
        (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | H=x) \<and>
        (is_quasilist(M,xs) | empty(M,H))" *)
 definition
-  hd_fm :: "[i,i]=>i" where
+  hd_fm :: "[i,i]\<Rightarrow>i" where
     "hd_fm(xs,H) \<equiv> 
        And(Implies(Nil_fm(xs), empty_fm(H)),
            And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(H#+2,1)))),
@@ -204,7 +204,7 @@
        (\<forall>x[M]. \<forall>l[M]. \<not> is_Cons(M,x,l,xs) | T=l) \<and>
        (is_quasilist(M,xs) | empty(M,T))" *)
 definition
-  tl_fm :: "[i,i]=>i" where
+  tl_fm :: "[i,i]\<Rightarrow>i" where
     "tl_fm(xs,T) \<equiv>
        And(Implies(Nil_fm(xs), Equal(T,xs)),
            And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))),
@@ -236,12 +236,12 @@
 
 subsubsection\<open>The Operator \<^term>\<open>is_bool_of_o\<close>\<close>
 
-(*   is_bool_of_o :: "[i=>o, o, i] => o"
+(*   is_bool_of_o :: "[i\<Rightarrow>o, o, i] \<Rightarrow> o"
    "is_bool_of_o(M,P,z) \<equiv> (P \<and> number1(M,z)) | (\<not>P \<and> empty(M,z))" *)
 
 text\<open>The formula \<^term>\<open>p\<close> has no free variables.\<close>
 definition
-  bool_of_o_fm :: "[i, i]=>i" where
+  bool_of_o_fm :: "[i, i]\<Rightarrow>i" where
   "bool_of_o_fm(p,z) \<equiv> 
     Or(And(p,number1_fm(z)),
        And(Neg(p),empty_fm(z)))"
@@ -279,12 +279,12 @@
 text\<open>The two arguments of \<^term>\<open>p\<close> are always 1, 0. Remember that
  \<^term>\<open>p\<close> will be enclosed by three quantifiers.\<close>
 
-(* is_lambda :: "[i=>o, i, [i,i]=>o, i] => o"
+(* is_lambda :: "[i\<Rightarrow>o, i, [i,i]\<Rightarrow>o, i] \<Rightarrow> o"
     "is_lambda(M, A, is_b, z) \<equiv> 
        \<forall>p[M]. p \<in> z \<longleftrightarrow>
         (\<exists>u[M]. \<exists>v[M]. u\<in>A \<and> pair(M,u,v,p) \<and> is_b(u,v))" *)
 definition
-  lambda_fm :: "[i, i, i]=>i" where
+  lambda_fm :: "[i, i, i]\<Rightarrow>i" where
   "lambda_fm(p,A,z) \<equiv> 
     Forall(Iff(Member(0,succ(z)),
             Exists(Exists(And(Member(1,A#+3),
@@ -324,7 +324,7 @@
 (*    "is_Member(M,x,y,Z) \<equiv>
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inl(M,p,u) \<and> is_Inl(M,u,Z)" *)
 definition
-  Member_fm :: "[i,i,i]=>i" where
+  Member_fm :: "[i,i,i]\<Rightarrow>i" where
     "Member_fm(x,y,Z) \<equiv>
        Exists(Exists(And(pair_fm(x#+2,y#+2,1), 
                       And(Inl_fm(1,0), Inl_fm(0,Z#+2)))))"
@@ -357,7 +357,7 @@
 (*    "is_Equal(M,x,y,Z) \<equiv>
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inr(M,p,u) \<and> is_Inl(M,u,Z)" *)
 definition
-  Equal_fm :: "[i,i,i]=>i" where
+  Equal_fm :: "[i,i,i]\<Rightarrow>i" where
     "Equal_fm(x,y,Z) \<equiv>
        Exists(Exists(And(pair_fm(x#+2,y#+2,1), 
                       And(Inr_fm(1,0), Inl_fm(0,Z#+2)))))"
@@ -390,7 +390,7 @@
 (*    "is_Nand(M,x,y,Z) \<equiv>
         \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) \<and> is_Inl(M,p,u) \<and> is_Inr(M,u,Z)" *)
 definition
-  Nand_fm :: "[i,i,i]=>i" where
+  Nand_fm :: "[i,i,i]\<Rightarrow>i" where
     "Nand_fm(x,y,Z) \<equiv>
        Exists(Exists(And(pair_fm(x#+2,y#+2,1), 
                       And(Inl_fm(1,0), Inr_fm(0,Z#+2)))))"
@@ -422,7 +422,7 @@
 
 (* "is_Forall(M,p,Z) \<equiv> \<exists>u[M]. is_Inr(M,p,u) \<and> is_Inr(M,u,Z)" *)
 definition
-  Forall_fm :: "[i,i]=>i" where
+  Forall_fm :: "[i,i]\<Rightarrow>i" where
     "Forall_fm(x,Z) \<equiv>
        Exists(And(Inr_fm(succ(x),0), Inr_fm(0,succ(Z))))"
 
@@ -455,7 +455,7 @@
 (* is_and(M,a,b,z) \<equiv> (number1(M,a)  \<and> z=b) | 
                        (\<not>number1(M,a) \<and> empty(M,z)) *)
 definition
-  and_fm :: "[i,i,i]=>i" where
+  and_fm :: "[i,i,i]\<Rightarrow>i" where
     "and_fm(a,b,z) \<equiv>
        Or(And(number1_fm(a), Equal(z,b)),
           And(Neg(number1_fm(a)),empty_fm(z)))"
@@ -490,7 +490,7 @@
                      (\<not>number1(M,a) \<and> z=b) *)
 
 definition
-  or_fm :: "[i,i,i]=>i" where
+  or_fm :: "[i,i,i]\<Rightarrow>i" where
     "or_fm(a,b,z) \<equiv>
        Or(And(number1_fm(a), number1_fm(z)),
           And(Neg(number1_fm(a)), Equal(z,b)))"
@@ -525,7 +525,7 @@
 (* is_not(M,a,z) \<equiv> (number1(M,a)  \<and> empty(M,z)) | 
                      (\<not>number1(M,a) \<and> number1(M,z)) *)
 definition
-  not_fm :: "[i,i]=>i" where
+  not_fm :: "[i,i]\<Rightarrow>i" where
     "not_fm(a,z) \<equiv>
        Or(And(number1_fm(a), empty_fm(z)),
           And(Neg(number1_fm(a)), number1_fm(z)))"
@@ -578,7 +578,7 @@
 done
 
 
-(* M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
+(* M_is_recfun :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i, i] \<Rightarrow> o"
    "M_is_recfun(M,MH,r,a,f) \<equiv>
      \<forall>z[M]. z \<in> f \<longleftrightarrow>
                2      1           0
@@ -592,7 +592,7 @@
 
 text\<open>The three arguments of \<^term>\<open>p\<close> are always 2, 1, 0 and z\<close>
 definition
-  is_recfun_fm :: "[i, i, i, i]=>i" where
+  is_recfun_fm :: "[i, i, i, i]\<Rightarrow>i" where
   "is_recfun_fm(p,r,a,f) \<equiv> 
    Forall(Iff(Member(0,succ(f)),
     Exists(Exists(Exists(
@@ -651,11 +651,11 @@
 text\<open>The three arguments of \<^term>\<open>p\<close> are always 2, 1, 0;
       \<^term>\<open>p\<close> is enclosed by 5 quantifiers.\<close>
 
-(* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o"
+(* is_wfrec :: "[i\<Rightarrow>o, i, [i,i,i]\<Rightarrow>o, i, i] \<Rightarrow> o"
     "is_wfrec(M,MH,r,a,z) \<equiv> 
       \<exists>f[M]. M_is_recfun(M,MH,r,a,f) \<and> MH(a,f,z)" *)
 definition
-  is_wfrec_fm :: "[i, i, i, i]=>i" where
+  is_wfrec_fm :: "[i, i, i, i]\<Rightarrow>i" where
   "is_wfrec_fm(p,r,a,z) \<equiv> 
     Exists(And(is_recfun_fm(p, succ(r), succ(a), 0),
            Exists(Exists(Exists(Exists(
@@ -714,7 +714,7 @@
 subsubsection\<open>Binary Products, Internalized\<close>
 
 definition
-  cartprod_fm :: "[i,i,i]=>i" where
+  cartprod_fm :: "[i,i,i]\<Rightarrow>i" where
 (* "cartprod(M,A,B,z) \<equiv>
         \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A \<and> (\<exists>y[M]. y\<in>B \<and> pair(M,x,y,u)))" *)
     "cartprod_fm(A,B,z) \<equiv>
@@ -755,7 +755,7 @@
        number1(M,n1) \<and> cartprod(M,n1,A,A0) \<and> upair(M,n1,n1,s1) \<and>
        cartprod(M,s1,B,B1) \<and> union(M,A0,B1,Z)"  *)
 definition
-  sum_fm :: "[i,i,i]=>i" where
+  sum_fm :: "[i,i,i]\<Rightarrow>i" where
     "sum_fm(A,B,Z) \<equiv>
        Exists(Exists(Exists(Exists(
         And(number1_fm(2),
@@ -791,7 +791,7 @@
 
 (* "is_quasinat(M,z) \<equiv> empty(M,z) | (\<exists>m[M]. successor(M,m,z))" *)
 definition
-  quasinat_fm :: "i=>i" where
+  quasinat_fm :: "i\<Rightarrow>i" where
     "quasinat_fm(z) \<equiv> Or(empty_fm(z), Exists(succ_fm(0,succ(z))))"
 
 lemma quasinat_type [TC]:
@@ -822,14 +822,14 @@
  \<^term>\<open>is_b\<close> takes two arguments.  Instead it must be a formula where 1 and 0
  stand for \<^term>\<open>m\<close> and \<^term>\<open>b\<close>, respectively.\<close>
 
-(* is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o"
+(* is_nat_case :: "[i\<Rightarrow>o, i, [i,i]\<Rightarrow>o, i, i] \<Rightarrow> o"
     "is_nat_case(M, a, is_b, k, z) \<equiv>
        (empty(M,k) \<longrightarrow> z=a) \<and>
        (\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) \<and>
        (is_quasinat(M,k) | empty(M,z))" *)
 text\<open>The formula \<^term>\<open>is_b\<close> has free variables 1 and 0.\<close>
 definition
-  is_nat_case_fm :: "[i, i, i, i]=>i" where
+  is_nat_case_fm :: "[i, i, i, i]\<Rightarrow>i" where
  "is_nat_case_fm(a,is_b,k,z) \<equiv> 
     And(Implies(empty_fm(k), Equal(z,a)),
         And(Forall(Implies(succ_fm(0,succ(k)), 
@@ -880,12 +880,12 @@
 
 subsection\<open>The Operator \<^term>\<open>iterates_MH\<close>, Needed for Iteration\<close>
 
-(*  iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o"
+(*  iterates_MH :: "[i\<Rightarrow>o, [i,i]\<Rightarrow>o, i, i, i, i] \<Rightarrow> o"
    "iterates_MH(M,isF,v,n,g,z) \<equiv>
         is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) \<and> isF(gm,u),
                     n, z)" *)
 definition
-  iterates_MH_fm :: "[i, i, i, i, i]=>i" where
+  iterates_MH_fm :: "[i, i, i, i, i]\<Rightarrow>i" where
  "iterates_MH_fm(isF,v,n,g,z) \<equiv> 
     is_nat_case_fm(v, 
       Exists(And(fun_apply_fm(succ(succ(succ(g))),2,0), 
@@ -951,7 +951,7 @@
        1       0       is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"*)
 
 definition
-  is_iterates_fm :: "[i, i, i, i]=>i" where
+  is_iterates_fm :: "[i, i, i, i]\<Rightarrow>i" where
   "is_iterates_fm(p,v,n,Z) \<equiv> 
      Exists(Exists(
       And(succ_fm(n#+2,1),
@@ -1022,7 +1022,7 @@
 (* is_eclose_n(M,A,n,Z) \<equiv> is_iterates(M, big_union(M), A, n, Z) *)
 
 definition
-  eclose_n_fm :: "[i,i,i]=>i" where
+  eclose_n_fm :: "[i,i,i]\<Rightarrow>i" where
   "eclose_n_fm(A,n,Z) \<equiv> is_iterates_fm(big_union_fm(1,0), A, n, Z)"
 
 lemma eclose_n_fm_type [TC]:
@@ -1059,7 +1059,7 @@
       \<exists>n[M]. \<exists>eclosen[M]. 
        finite_ordinal(M,n) \<and> is_eclose_n(M,A,n,eclosen) \<and> l \<in> eclosen *)
 definition
-  mem_eclose_fm :: "[i,i]=>i" where
+  mem_eclose_fm :: "[i,i]\<Rightarrow>i" where
     "mem_eclose_fm(x,y) \<equiv>
        Exists(Exists(
          And(finite_ordinal_fm(1),
@@ -1092,7 +1092,7 @@
 
 (* is_eclose(M,A,Z) \<equiv> \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_eclose(M,A,l) *)
 definition
-  is_eclose_fm :: "[i,i]=>i" where
+  is_eclose_fm :: "[i,i]\<Rightarrow>i" where
     "is_eclose_fm(A,Z) \<equiv>
        Forall(Iff(Member(0,succ(Z)), mem_eclose_fm(succ(A),0)))"
 
@@ -1122,7 +1122,7 @@
 subsubsection\<open>The List Functor, Internalized\<close>
 
 definition
-  list_functor_fm :: "[i,i,i]=>i" where
+  list_functor_fm :: "[i,i,i]\<Rightarrow>i" where
 (* "is_list_functor(M,A,X,Z) \<equiv>
         \<exists>n1[M]. \<exists>AX[M].
          number1(M,n1) \<and> cartprod(M,A,X,AX) \<and> is_sum(M,n1,AX,Z)" *)
@@ -1163,7 +1163,7 @@
                 is_iterates(M, is_list_functor(M,A), zero, n, Z)" *)
 
 definition
-  list_N_fm :: "[i,i,i]=>i" where
+  list_N_fm :: "[i,i,i]\<Rightarrow>i" where
   "list_N_fm(A,n,Z) \<equiv> 
      Exists(
        And(empty_fm(0),
@@ -1204,7 +1204,7 @@
       \<exists>n[M]. \<exists>listn[M]. 
        finite_ordinal(M,n) \<and> is_list_N(M,A,n,listn) \<and> l \<in> listn *)
 definition
-  mem_list_fm :: "[i,i]=>i" where
+  mem_list_fm :: "[i,i]\<Rightarrow>i" where
     "mem_list_fm(x,y) \<equiv>
        Exists(Exists(
          And(finite_ordinal_fm(1),
@@ -1237,7 +1237,7 @@
 
 (* is_list(M,A,Z) \<equiv> \<forall>l[M]. l \<in> Z \<longleftrightarrow> mem_list(M,A,l) *)
 definition
-  is_list_fm :: "[i,i]=>i" where
+  is_list_fm :: "[i,i]\<Rightarrow>i" where
     "is_list_fm(A,Z) \<equiv>
        Forall(Iff(Member(0,succ(Z)), mem_list_fm(succ(A),0)))"
 
@@ -1266,7 +1266,7 @@
 
 subsubsection\<open>The Formula Functor, Internalized\<close>
 
-definition formula_functor_fm :: "[i,i]=>i" where
+definition formula_functor_fm :: "[i,i]\<Rightarrow>i" where
 (*     "is_formula_functor(M,X,Z) \<equiv>
         \<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M].
            4           3               2       1       0
@@ -1313,7 +1313,7 @@
       \<exists>zero[M]. empty(M,zero) \<and> 
                 is_iterates(M, is_formula_functor(M), zero, n, Z)" *) 
 definition
-  formula_N_fm :: "[i,i]=>i" where
+  formula_N_fm :: "[i,i]\<Rightarrow>i" where
   "formula_N_fm(n,Z) \<equiv> 
      Exists(
        And(empty_fm(0),
@@ -1354,7 +1354,7 @@
       \<exists>n[M]. \<exists>formn[M]. 
        finite_ordinal(M,n) \<and> is_formula_N(M,n,formn) \<and> p \<in> formn *)
 definition
-  mem_formula_fm :: "i=>i" where
+  mem_formula_fm :: "i\<Rightarrow>i" where
     "mem_formula_fm(x) \<equiv>
        Exists(Exists(
          And(finite_ordinal_fm(1),
@@ -1387,7 +1387,7 @@
 
 (* is_formula(M,Z) \<equiv> \<forall>p[M]. p \<in> Z \<longleftrightarrow> mem_formula(M,p) *)
 definition
-  is_formula_fm :: "i=>i" where
+  is_formula_fm :: "i\<Rightarrow>i" where
     "is_formula_fm(Z) \<equiv> Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))"
 
 lemma is_formula_type [TC]:
@@ -1419,14 +1419,14 @@
    We call \<^term>\<open>p\<close> with arguments a, f, z by equating them with 
   the corresponding quantified variables with de Bruijn indices 2, 1, 0.\<close>
 
-(* is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o"
+(* is_transrec :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i] \<Rightarrow> o"
    "is_transrec(M,MH,a,z) \<equiv> 
       \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M]. 
        2       1         0
        upair(M,a,a,sa) \<and> is_eclose(M,sa,esa) \<and> membership(M,esa,mesa) \<and>
        is_wfrec(M,MH,mesa,a,z)" *)
 definition
-  is_transrec_fm :: "[i, i, i]=>i" where
+  is_transrec_fm :: "[i, i, i]\<Rightarrow>i" where
  "is_transrec_fm(p,a,z) \<equiv> 
     Exists(Exists(Exists(
       And(upair_fm(a#+3,a#+3,2),
--- a/src/ZF/Constructible/L_axioms.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/L_axioms.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -53,8 +53,8 @@
 
 lemma LReplace_in_Lset:
      "\<lbrakk>X \<in> Lset(i); univalent(L,X,Q); Ord(i)\<rbrakk>
-      \<Longrightarrow> \<exists>j. Ord(j) \<and> Replace(X, %x y. Q(x,y) \<and> L(y)) \<subseteq> Lset(j)"
-apply (rule_tac x="\<Union>y \<in> Replace(X, %x y. Q(x,y) \<and> L(y)). succ(lrank(y))"
+      \<Longrightarrow> \<exists>j. Ord(j) \<and> Replace(X, \<lambda>x y. Q(x,y) \<and> L(y)) \<subseteq> Lset(j)"
+apply (rule_tac x="\<Union>y \<in> Replace(X, \<lambda>x y. Q(x,y) \<and> L(y)). succ(lrank(y))"
        in exI)
 apply simp
 apply clarify
@@ -65,7 +65,7 @@
 
 lemma LReplace_in_L:
      "\<lbrakk>L(X); univalent(L,X,Q)\<rbrakk>
-      \<Longrightarrow> \<exists>Y. L(Y) \<and> Replace(X, %x y. Q(x,y) \<and> L(y)) \<subseteq> Y"
+      \<Longrightarrow> \<exists>Y. L(Y) \<and> Replace(X, \<lambda>x y. Q(x,y) \<and> L(y)) \<subseteq> Y"
 apply (drule L_D, clarify)
 apply (drule LReplace_in_Lset, assumption+)
 apply (blast intro: L_I Lset_in_Lset_succ)
@@ -79,7 +79,7 @@
 done
 
 lemma strong_replacementI [rule_format]:
-    "\<lbrakk>\<forall>B[L]. separation(L, %u. \<exists>x[L]. x\<in>B \<and> P(x,u))\<rbrakk>
+    "\<lbrakk>\<forall>B[L]. separation(L, \<lambda>u. \<exists>x[L]. x\<in>B \<and> P(x,u))\<rbrakk>
      \<Longrightarrow> strong_replacement(L,P)"
 apply (simp add: strong_replacement_def, clarify)
 apply (frule replacementD [OF replacement], assumption, clarify)
@@ -126,22 +126,22 @@
 text\<open>instances of locale constants\<close>
 
 definition
-  L_F0 :: "[i=>o,i] => i" where
-    "L_F0(P,y) \<equiv> \<mu> b. (\<exists>z. L(z) \<and> P(<y,z>)) \<longrightarrow> (\<exists>z\<in>Lset(b). P(<y,z>))"
+  L_F0 :: "[i\<Rightarrow>o,i] \<Rightarrow> i" where
+    "L_F0(P,y) \<equiv> \<mu> b. (\<exists>z. L(z) \<and> P(\<langle>y,z\<rangle>)) \<longrightarrow> (\<exists>z\<in>Lset(b). P(\<langle>y,z\<rangle>))"
 
 definition
-  L_FF :: "[i=>o,i] => i" where
+  L_FF :: "[i\<Rightarrow>o,i] \<Rightarrow> i" where
     "L_FF(P)   \<equiv> \<lambda>a. \<Union>y\<in>Lset(a). L_F0(P,y)"
 
 definition
-  L_ClEx :: "[i=>o,i] => o" where
+  L_ClEx :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "L_ClEx(P) \<equiv> \<lambda>a. Limit(a) \<and> normalize(L_FF(P),a) = a"
 
 
 text\<open>We must use the meta-existential quantifier; otherwise the reflection
       terms become enormous!\<close>
 definition
-  L_Reflects :: "[i=>o,[i,i]=>o] => prop"  (\<open>(3REFLECTS/ [_,/ _])\<close>) where
+  L_Reflects :: "[i\<Rightarrow>o,[i,i]\<Rightarrow>o] \<Rightarrow> prop"  (\<open>(3REFLECTS/ [_,/ _])\<close>) where
     "REFLECTS[P,Q] \<equiv> (\<Or>Cl. Closed_Unbounded(Cl) \<and>
                            (\<forall>a. Cl(a) \<longrightarrow> (\<forall>x \<in> Lset(a). P(x) \<longleftrightarrow> Q(a,x))))"
 
@@ -303,7 +303,7 @@
 subsubsection\<open>The Empty Set, Internalized\<close>
 
 definition
-  empty_fm :: "i=>i" where
+  empty_fm :: "i\<Rightarrow>i" where
     "empty_fm(x) \<equiv> Forall(Neg(Member(0,succ(x))))"
 
 lemma empty_type [TC]:
@@ -342,7 +342,7 @@
 subsubsection\<open>Unordered Pairs, Internalized\<close>
 
 definition
-  upair_fm :: "[i,i,i]=>i" where
+  upair_fm :: "[i,i,i]\<Rightarrow>i" where
     "upair_fm(x,y,z) \<equiv>
        And(Member(x,z),
            And(Member(y,z),
@@ -385,7 +385,7 @@
 subsubsection\<open>Ordered pairs, Internalized\<close>
 
 definition
-  pair_fm :: "[i,i,i]=>i" where
+  pair_fm :: "[i,i,i]\<Rightarrow>i" where
     "pair_fm(x,y,z) \<equiv>
        Exists(And(upair_fm(succ(x),succ(x),0),
               Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0),
@@ -418,7 +418,7 @@
 subsubsection\<open>Binary Unions, Internalized\<close>
 
 definition
-  union_fm :: "[i,i,i]=>i" where
+  union_fm :: "[i,i,i]\<Rightarrow>i" where
     "union_fm(x,y,z) \<equiv>
        Forall(Iff(Member(0,succ(z)),
                   Or(Member(0,succ(x)),Member(0,succ(y)))))"
@@ -450,7 +450,7 @@
 subsubsection\<open>Set ``Cons,'' Internalized\<close>
 
 definition
-  cons_fm :: "[i,i,i]=>i" where
+  cons_fm :: "[i,i,i]\<Rightarrow>i" where
     "cons_fm(x,y,z) \<equiv>
        Exists(And(upair_fm(succ(x),succ(x),0),
                   union_fm(0,succ(y),succ(z))))"
@@ -483,7 +483,7 @@
 subsubsection\<open>Successor Function, Internalized\<close>
 
 definition
-  succ_fm :: "[i,i]=>i" where
+  succ_fm :: "[i,i]\<Rightarrow>i" where
     "succ_fm(x,y) \<equiv> cons_fm(x,x,y)"
 
 lemma succ_type [TC]:
@@ -514,7 +514,7 @@
 
 (* "number1(M,a) \<equiv> (\<exists>x[M]. empty(M,x) \<and> successor(M,x,a))" *)
 definition
-  number1_fm :: "i=>i" where
+  number1_fm :: "i\<Rightarrow>i" where
     "number1_fm(a) \<equiv> Exists(And(empty_fm(0), succ_fm(0,succ(a))))"
 
 lemma number1_type [TC]:
@@ -544,7 +544,7 @@
 
 (*  "big_union(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A \<and> x \<in> y)" *)
 definition
-  big_union_fm :: "[i,i]=>i" where
+  big_union_fm :: "[i,i]\<Rightarrow>i" where
     "big_union_fm(A,z) \<equiv>
        Forall(Iff(Member(0,succ(z)),
                   Exists(And(Member(0,succ(succ(A))), Member(1,0)))))"
@@ -625,7 +625,7 @@
 subsubsection\<open>Membership Relation, Internalized\<close>
 
 definition
-  Memrel_fm :: "[i,i]=>i" where
+  Memrel_fm :: "[i,i]\<Rightarrow>i" where
     "Memrel_fm(A,r) \<equiv>
        Forall(Iff(Member(0,succ(r)),
                   Exists(And(Member(0,succ(succ(A))),
@@ -659,7 +659,7 @@
 subsubsection\<open>Predecessor Set, Internalized\<close>
 
 definition
-  pred_set_fm :: "[i,i,i,i]=>i" where
+  pred_set_fm :: "[i,i,i,i]\<Rightarrow>i" where
     "pred_set_fm(A,x,r,B) \<equiv>
        Forall(Iff(Member(0,succ(B)),
                   Exists(And(Member(0,succ(succ(r))),
@@ -698,7 +698,7 @@
 (* "is_domain(M,r,z) \<equiv>
         \<forall>x[M]. (x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>y[M]. pair(M,x,y,w))))" *)
 definition
-  domain_fm :: "[i,i]=>i" where
+  domain_fm :: "[i,i]\<Rightarrow>i" where
     "domain_fm(r,z) \<equiv>
        Forall(Iff(Member(0,succ(z)),
                   Exists(And(Member(0,succ(succ(r))),
@@ -733,7 +733,7 @@
 (* "is_range(M,r,z) \<equiv>
         \<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>x[M]. pair(M,x,y,w))))" *)
 definition
-  range_fm :: "[i,i]=>i" where
+  range_fm :: "[i,i]\<Rightarrow>i" where
     "range_fm(r,z) \<equiv>
        Forall(Iff(Member(0,succ(z)),
                   Exists(And(Member(0,succ(succ(r))),
@@ -769,7 +769,7 @@
         \<exists>dr[M]. is_domain(M,r,dr) \<and>
             (\<exists>rr[M]. is_range(M,r,rr) \<and> union(M,dr,rr,z))" *)
 definition
-  field_fm :: "[i,i]=>i" where
+  field_fm :: "[i,i]\<Rightarrow>i" where
     "field_fm(r,z) \<equiv>
        Exists(And(domain_fm(succ(r),0),
               Exists(And(range_fm(succ(succ(r)),0),
@@ -805,7 +805,7 @@
 (* "image(M,r,A,z) \<equiv>
         \<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>x[M]. x\<in>A \<and> pair(M,x,y,w))))" *)
 definition
-  image_fm :: "[i,i,i]=>i" where
+  image_fm :: "[i,i,i]\<Rightarrow>i" where
     "image_fm(r,A,z) \<equiv>
        Forall(Iff(Member(0,succ(z)),
                   Exists(And(Member(0,succ(succ(r))),
@@ -841,7 +841,7 @@
 (* "pre_image(M,r,A,z) \<equiv>
         \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>y[M]. y\<in>A \<and> pair(M,x,y,w)))" *)
 definition
-  pre_image_fm :: "[i,i,i]=>i" where
+  pre_image_fm :: "[i,i,i]\<Rightarrow>i" where
     "pre_image_fm(r,A,z) \<equiv>
        Forall(Iff(Member(0,succ(z)),
                   Exists(And(Member(0,succ(succ(r))),
@@ -878,7 +878,7 @@
         (\<exists>xs[M]. \<exists>fxs[M].
          upair(M,x,x,xs) \<and> image(M,f,xs,fxs) \<and> big_union(M,fxs,y))" *)
 definition
-  fun_apply_fm :: "[i,i,i]=>i" where
+  fun_apply_fm :: "[i,i,i]\<Rightarrow>i" where
     "fun_apply_fm(f,x,y) \<equiv>
        Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1),
                          And(image_fm(succ(succ(f)), 1, 0),
@@ -914,7 +914,7 @@
 (* "is_relation(M,r) \<equiv>
         (\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *)
 definition
-  relation_fm :: "i=>i" where
+  relation_fm :: "i\<Rightarrow>i" where
     "relation_fm(r) \<equiv>
        Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))"
 
@@ -947,7 +947,7 @@
         \<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
            pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'" *)
 definition
-  function_fm :: "i=>i" where
+  function_fm :: "i\<Rightarrow>i" where
     "function_fm(r) \<equiv>
        Forall(Forall(Forall(Forall(Forall(
          Implies(pair_fm(4,3,1),
@@ -985,7 +985,7 @@
         (\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))" *)
 
 definition
-  typed_function_fm :: "[i,i,i]=>i" where
+  typed_function_fm :: "[i,i,i]\<Rightarrow>i" where
     "typed_function_fm(A,B,r) \<equiv>
        And(function_fm(r),
          And(relation_fm(r),
@@ -1045,7 +1045,7 @@
                 pair(M,x,z,p) \<and> pair(M,x,y,xy) \<and> pair(M,y,z,yz) \<and>
                 xy \<in> s \<and> yz \<in> r)" *)
 definition
-  composition_fm :: "[i,i,i]=>i" where
+  composition_fm :: "[i,i,i]\<Rightarrow>i" where
   "composition_fm(r,s,t) \<equiv>
      Forall(Iff(Member(0,succ(t)),
              Exists(Exists(Exists(Exists(Exists(
@@ -1085,7 +1085,7 @@
         (\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
           pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')" *)
 definition
-  injection_fm :: "[i,i,i]=>i" where
+  injection_fm :: "[i,i,i]\<Rightarrow>i" where
   "injection_fm(A,B,f) \<equiv>
     And(typed_function_fm(A,B,f),
        Forall(Forall(Forall(Forall(Forall(
@@ -1121,12 +1121,12 @@
 
 subsubsection\<open>Surjections, Internalized\<close>
 
-(*  surjection :: "[i=>o,i,i,i] => o"
+(*  surjection :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o"
     "surjection(M,A,B,f) \<equiv>
         typed_function(M,A,B,f) \<and>
         (\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A \<and> fun_apply(M,f,x,y)))" *)
 definition
-  surjection_fm :: "[i,i,i]=>i" where
+  surjection_fm :: "[i,i,i]\<Rightarrow>i" where
   "surjection_fm(A,B,f) \<equiv>
     And(typed_function_fm(A,B,f),
        Forall(Implies(Member(0,succ(B)),
@@ -1160,10 +1160,10 @@
 
 subsubsection\<open>Bijections, Internalized\<close>
 
-(*   bijection :: "[i=>o,i,i,i] => o"
+(*   bijection :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o"
     "bijection(M,A,B,f) \<equiv> injection(M,A,B,f) \<and> surjection(M,A,B,f)" *)
 definition
-  bijection_fm :: "[i,i,i]=>i" where
+  bijection_fm :: "[i,i,i]\<Rightarrow>i" where
   "bijection_fm(A,B,f) \<equiv> And(injection_fm(A,B,f), surjection_fm(A,B,f))"
 
 lemma bijection_type [TC]:
@@ -1196,7 +1196,7 @@
 (* "restriction(M,r,A,z) \<equiv>
         \<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r \<and> (\<exists>u[M]. u\<in>A \<and> (\<exists>v[M]. pair(M,u,v,x))))" *)
 definition
-  restriction_fm :: "[i,i,i]=>i" where
+  restriction_fm :: "[i,i,i]\<Rightarrow>i" where
     "restriction_fm(r,A,z) \<equiv>
        Forall(Iff(Member(0,succ(z)),
                   And(Member(0,succ(r)),
@@ -1228,7 +1228,7 @@
 
 subsubsection\<open>Order-Isomorphisms, Internalized\<close>
 
-(*  order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
+(*  order_isomorphism :: "[i\<Rightarrow>o,i,i,i,i,i] \<Rightarrow> o"
    "order_isomorphism(M,A,r,B,s,f) \<equiv>
         bijection(M,A,B,f) \<and>
         (\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
@@ -1238,7 +1238,7 @@
   *)
 
 definition
-  order_isomorphism_fm :: "[i,i,i,i,i]=>i" where
+  order_isomorphism_fm :: "[i,i,i,i,i]\<Rightarrow>i" where
  "order_isomorphism_fm(A,r,B,s,f) \<equiv>
    And(bijection_fm(A,B,f),
      Forall(Implies(Member(0,succ(A)),
@@ -1286,7 +1286,7 @@
         (\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a \<and> successor(M,x,y)))" *)
 
 definition
-  limit_ordinal_fm :: "i=>i" where
+  limit_ordinal_fm :: "i\<Rightarrow>i" where
     "limit_ordinal_fm(x) \<equiv>
         And(ordinal_fm(x),
             And(Neg(empty_fm(x)),
@@ -1323,7 +1323,7 @@
         ordinal(M,a) \<and> \<not> limit_ordinal(M,a) \<and> 
         (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))" *)
 definition
-  finite_ordinal_fm :: "i=>i" where
+  finite_ordinal_fm :: "i\<Rightarrow>i" where
     "finite_ordinal_fm(x) \<equiv>
        And(ordinal_fm(x),
           And(Neg(limit_ordinal_fm(x)),
@@ -1357,7 +1357,7 @@
 
 (* omega(M,a) \<equiv> limit_ordinal(M,a) \<and> (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x)) *)
 definition
-  omega_fm :: "i=>i" where
+  omega_fm :: "i\<Rightarrow>i" where
     "omega_fm(x) \<equiv>
        And(limit_ordinal_fm(x),
            Forall(Implies(Member(0,succ(x)),
--- a/src/ZF/Constructible/Normal.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Normal.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -18,60 +18,60 @@
 subsection \<open>Closed and Unbounded (c.u.) Classes of Ordinals\<close>
 
 definition
-  Closed :: "(i=>o) => o" where
+  Closed :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "Closed(P) \<equiv> \<forall>I. I \<noteq> 0 \<longrightarrow> (\<forall>i\<in>I. Ord(i) \<and> P(i)) \<longrightarrow> P(\<Union>(I))"
 
 definition
-  Unbounded :: "(i=>o) => o" where
+  Unbounded :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "Unbounded(P) \<equiv> \<forall>i. Ord(i) \<longrightarrow> (\<exists>j. i<j \<and> P(j))"
 
 definition
-  Closed_Unbounded :: "(i=>o) => o" where
+  Closed_Unbounded :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "Closed_Unbounded(P) \<equiv> Closed(P) \<and> Unbounded(P)"
 
 
 subsubsection\<open>Simple facts about c.u. classes\<close>
 
 lemma ClosedI:
-     "\<lbrakk>\<And>I. \<lbrakk>I \<noteq> 0; \<forall>i\<in>I. Ord(i) \<and> P(i)\<rbrakk> \<Longrightarrow> P(\<Union>(I))\<rbrakk> 
+  "\<lbrakk>\<And>I. \<lbrakk>I \<noteq> 0; \<forall>i\<in>I. Ord(i) \<and> P(i)\<rbrakk> \<Longrightarrow> P(\<Union>(I))\<rbrakk> 
       \<Longrightarrow> Closed(P)"
-by (simp add: Closed_def)
+  by (simp add: Closed_def)
 
 lemma ClosedD:
-     "\<lbrakk>Closed(P); I \<noteq> 0; \<And>i. i\<in>I \<Longrightarrow> Ord(i); \<And>i. i\<in>I \<Longrightarrow> P(i)\<rbrakk> 
+  "\<lbrakk>Closed(P); I \<noteq> 0; \<And>i. i\<in>I \<Longrightarrow> Ord(i); \<And>i. i\<in>I \<Longrightarrow> P(i)\<rbrakk> 
       \<Longrightarrow> P(\<Union>(I))"
-by (simp add: Closed_def)
+  by (simp add: Closed_def)
 
 lemma UnboundedD:
-     "\<lbrakk>Unbounded(P);  Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. i<j \<and> P(j)"
-by (simp add: Unbounded_def)
+  "\<lbrakk>Unbounded(P);  Ord(i)\<rbrakk> \<Longrightarrow> \<exists>j. i<j \<and> P(j)"
+  by (simp add: Unbounded_def)
 
 lemma Closed_Unbounded_imp_Unbounded: "Closed_Unbounded(C) \<Longrightarrow> Unbounded(C)"
-by (simp add: Closed_Unbounded_def) 
+  by (simp add: Closed_Unbounded_def) 
 
 
 text\<open>The universal class, V, is closed and unbounded.
       A bit odd, since C. U. concerns only ordinals, but it's used below!\<close>
 theorem Closed_Unbounded_V [simp]: "Closed_Unbounded(\<lambda>x. True)"
-by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast)
+  by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast)
 
 text\<open>The class of ordinals, \<^term>\<open>Ord\<close>, is closed and unbounded.\<close>
 theorem Closed_Unbounded_Ord   [simp]: "Closed_Unbounded(Ord)"
-by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast)
+  by (unfold Closed_Unbounded_def Closed_def Unbounded_def, blast)
 
 text\<open>The class of limit ordinals, \<^term>\<open>Limit\<close>, is closed and unbounded.\<close>
 theorem Closed_Unbounded_Limit [simp]: "Closed_Unbounded(Limit)"
-apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def Limit_Union, 
-       clarify)
-apply (rule_tac x="i++nat" in exI)  
-apply (blast intro: oadd_lt_self oadd_LimitI Limit_has_0) 
-done
+  apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def Limit_Union, 
+      clarify)
+  apply (rule_tac x="i++nat" in exI)  
+  apply (blast intro: oadd_lt_self oadd_LimitI Limit_has_0) 
+  done
 
 text\<open>The class of cardinals, \<^term>\<open>Card\<close>, is closed and unbounded.\<close>
 theorem Closed_Unbounded_Card  [simp]: "Closed_Unbounded(Card)"
-apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def)
-apply (blast intro: lt_csucc Card_csucc)
-done
+  apply (simp add: Closed_Unbounded_def Closed_def Unbounded_def)
+  apply (blast intro: lt_csucc Card_csucc)
+  done
 
 
 subsubsection\<open>The intersection of any set-indexed family of c.u. classes is
@@ -87,160 +87,163 @@
       and A_non0: "A\<noteq>0"
   defines "next_greater(a,x) \<equiv> \<mu> y. x<y \<and> P(a,y)"
       and "sup_greater(x) \<equiv> \<Union>a\<in>A. next_greater(a,x)"
- 
+
+begin
 
 text\<open>Trivial that the intersection is closed.\<close>
-lemma (in cub_family) Closed_INT: "Closed(\<lambda>x. \<forall>i\<in>A. P(i,x))"
-by (blast intro: ClosedI ClosedD [OF closed])
+lemma Closed_INT: "Closed(\<lambda>x. \<forall>i\<in>A. P(i,x))"
+  by (blast intro: ClosedI ClosedD [OF closed])
 
 text\<open>All remaining effort goes to show that the intersection is unbounded.\<close>
 
-lemma (in cub_family) Ord_sup_greater:
-     "Ord(sup_greater(x))"
-by (simp add: sup_greater_def next_greater_def)
+lemma Ord_sup_greater:
+  "Ord(sup_greater(x))"
+  by (simp add: sup_greater_def next_greater_def)
 
-lemma (in cub_family) Ord_next_greater:
-     "Ord(next_greater(a,x))"
-by (simp add: next_greater_def)
+lemma Ord_next_greater:
+  "Ord(next_greater(a,x))"
+  by (simp add: next_greater_def)
 
 text\<open>\<^term>\<open>next_greater\<close> works as expected: it returns a larger value
 and one that belongs to class \<^term>\<open>P(a)\<close>.\<close>
-lemma (in cub_family) next_greater_lemma:
-     "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, next_greater(a,x)) \<and> x < next_greater(a,x)"
-apply (simp add: next_greater_def)
-apply (rule exE [OF UnboundedD [OF unbounded]])
-  apply assumption+
-apply (blast intro: LeastI2 lt_Ord2) 
-done
+lemma next_greater_lemma:
+  "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, next_greater(a,x)) \<and> x < next_greater(a,x)"
+  apply (simp add: next_greater_def)
+  apply (rule exE [OF UnboundedD [OF unbounded]])
+    apply assumption+
+  apply (blast intro: LeastI2 lt_Ord2) 
+  done
 
-lemma (in cub_family) next_greater_in_P:
-     "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, next_greater(a,x))"
-by (blast dest: next_greater_lemma)
+lemma next_greater_in_P:
+  "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, next_greater(a,x))"
+  by (blast dest: next_greater_lemma)
 
-lemma (in cub_family) next_greater_gt:
-     "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> x < next_greater(a,x)"
-by (blast dest: next_greater_lemma)
+lemma next_greater_gt:
+  "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> x < next_greater(a,x)"
+  by (blast dest: next_greater_lemma)
 
-lemma (in cub_family) sup_greater_gt:
-     "Ord(x) \<Longrightarrow> x < sup_greater(x)"
-apply (simp add: sup_greater_def)
-apply (insert A_non0)
-apply (blast intro: UN_upper_lt next_greater_gt Ord_next_greater)
-done
+lemma sup_greater_gt:
+  "Ord(x) \<Longrightarrow> x < sup_greater(x)"
+  apply (simp add: sup_greater_def)
+  apply (insert A_non0)
+  apply (blast intro: UN_upper_lt next_greater_gt Ord_next_greater)
+  done
 
-lemma (in cub_family) next_greater_le_sup_greater:
-     "a\<in>A \<Longrightarrow> next_greater(a,x) \<le> sup_greater(x)"
-apply (simp add: sup_greater_def) 
-apply (blast intro: UN_upper_le Ord_next_greater)
-done
+lemma next_greater_le_sup_greater:
+  "a\<in>A \<Longrightarrow> next_greater(a,x) \<le> sup_greater(x)"
+  apply (simp add: sup_greater_def) 
+  apply (blast intro: UN_upper_le Ord_next_greater)
+  done
 
-lemma (in cub_family) omega_sup_greater_eq_UN:
-     "\<lbrakk>Ord(x); a\<in>A\<rbrakk> 
+lemma omega_sup_greater_eq_UN:
+  "\<lbrakk>Ord(x); a\<in>A\<rbrakk> 
       \<Longrightarrow> sup_greater^\<omega> (x) = 
           (\<Union>n\<in>nat. next_greater(a, sup_greater^n (x)))"
-apply (simp add: iterates_omega_def)
-apply (rule le_anti_sym)
-apply (rule le_implies_UN_le_UN) 
-apply (blast intro: leI next_greater_gt Ord_iterates Ord_sup_greater)  
-txt\<open>Opposite bound:
+  apply (simp add: iterates_omega_def)
+  apply (rule le_anti_sym)
+   apply (rule le_implies_UN_le_UN) 
+   apply (blast intro: leI next_greater_gt Ord_iterates Ord_sup_greater)  
+  txt\<open>Opposite bound:
 @{subgoals[display,indent=0,margin=65]}
 \<close>
-apply (rule UN_least_le) 
-apply (blast intro: Ord_iterates Ord_sup_greater)  
-apply (rule_tac a="succ(n)" in UN_upper_le)
-apply (simp_all add: next_greater_le_sup_greater) 
-apply (blast intro: Ord_iterates Ord_sup_greater)  
-done
+  apply (rule UN_least_le) 
+   apply (blast intro: Ord_iterates Ord_sup_greater)  
+  apply (rule_tac a="succ(n)" in UN_upper_le)
+    apply (simp_all add: next_greater_le_sup_greater) 
+  apply (blast intro: Ord_iterates Ord_sup_greater)  
+  done
 
-lemma (in cub_family) P_omega_sup_greater:
-     "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, sup_greater^\<omega> (x))"
-apply (simp add: omega_sup_greater_eq_UN)
-apply (rule ClosedD [OF closed]) 
-apply (blast intro: ltD, auto)
-apply (blast intro: Ord_iterates Ord_next_greater Ord_sup_greater)
-apply (blast intro: next_greater_in_P Ord_iterates Ord_sup_greater)
-done
+lemma P_omega_sup_greater:
+  "\<lbrakk>Ord(x); a\<in>A\<rbrakk> \<Longrightarrow> P(a, sup_greater^\<omega> (x))"
+  apply (simp add: omega_sup_greater_eq_UN)
+  apply (rule ClosedD [OF closed]) 
+     apply (blast intro: ltD, auto)
+   apply (blast intro: Ord_iterates Ord_next_greater Ord_sup_greater)
+  apply (blast intro: next_greater_in_P Ord_iterates Ord_sup_greater)
+  done
 
-lemma (in cub_family) omega_sup_greater_gt:
-     "Ord(x) \<Longrightarrow> x < sup_greater^\<omega> (x)"
-apply (simp add: iterates_omega_def)
-apply (rule UN_upper_lt [of 1], simp_all) 
- apply (blast intro: sup_greater_gt) 
-apply (blast intro: Ord_iterates Ord_sup_greater)
-done
+lemma omega_sup_greater_gt:
+  "Ord(x) \<Longrightarrow> x < sup_greater^\<omega> (x)"
+  apply (simp add: iterates_omega_def)
+  apply (rule UN_upper_lt [of 1], simp_all) 
+   apply (blast intro: sup_greater_gt) 
+  apply (blast intro: Ord_iterates Ord_sup_greater)
+  done
 
-lemma (in cub_family) Unbounded_INT: "Unbounded(\<lambda>x. \<forall>a\<in>A. P(a,x))"
-apply (unfold Unbounded_def)  
-apply (blast intro!: omega_sup_greater_gt P_omega_sup_greater) 
-done
+lemma Unbounded_INT: "Unbounded(\<lambda>x. \<forall>a\<in>A. P(a,x))"
+  apply (unfold Unbounded_def)  
+  apply (blast intro!: omega_sup_greater_gt P_omega_sup_greater) 
+  done
 
-lemma (in cub_family) Closed_Unbounded_INT: 
-     "Closed_Unbounded(\<lambda>x. \<forall>a\<in>A. P(a,x))"
-by (simp add: Closed_Unbounded_def Closed_INT Unbounded_INT)
+lemma Closed_Unbounded_INT: 
+  "Closed_Unbounded(\<lambda>x. \<forall>a\<in>A. P(a,x))"
+  by (simp add: Closed_Unbounded_def Closed_INT Unbounded_INT)
+
+end
 
 
 theorem Closed_Unbounded_INT:
-    "(\<And>a. a\<in>A \<Longrightarrow> Closed_Unbounded(P(a)))
+  "(\<And>a. a\<in>A \<Longrightarrow> Closed_Unbounded(P(a)))
      \<Longrightarrow> Closed_Unbounded(\<lambda>x. \<forall>a\<in>A. P(a, x))"
-apply (case_tac "A=0", simp)
-apply (rule cub_family.Closed_Unbounded_INT [OF cub_family.intro])
-apply (simp_all add: Closed_Unbounded_def)
-done
+  apply (case_tac "A=0", simp)
+  apply (rule cub_family.Closed_Unbounded_INT [OF cub_family.intro])
+    apply (simp_all add: Closed_Unbounded_def)
+  done
 
 lemma Int_iff_INT2:
-     "P(x) \<and> Q(x)  \<longleftrightarrow>  (\<forall>i\<in>2. (i=0 \<longrightarrow> P(x)) \<and> (i=1 \<longrightarrow> Q(x)))"
-by auto
+  "P(x) \<and> Q(x)  \<longleftrightarrow>  (\<forall>i\<in>2. (i=0 \<longrightarrow> P(x)) \<and> (i=1 \<longrightarrow> Q(x)))"
+  by auto
 
 theorem Closed_Unbounded_Int:
-     "\<lbrakk>Closed_Unbounded(P); Closed_Unbounded(Q)\<rbrakk> 
+  "\<lbrakk>Closed_Unbounded(P); Closed_Unbounded(Q)\<rbrakk> 
       \<Longrightarrow> Closed_Unbounded(\<lambda>x. P(x) \<and> Q(x))"
-apply (simp only: Int_iff_INT2)
-apply (rule Closed_Unbounded_INT, auto) 
-done
+  apply (simp only: Int_iff_INT2)
+  apply (rule Closed_Unbounded_INT, auto) 
+  done
 
 
 subsection \<open>Normal Functions\<close> 
 
 definition
-  mono_le_subset :: "(i=>i) => o" where
+  mono_le_subset :: "(i\<Rightarrow>i) \<Rightarrow> o" where
     "mono_le_subset(M) \<equiv> \<forall>i j. i\<le>j \<longrightarrow> M(i) \<subseteq> M(j)"
 
 definition
-  mono_Ord :: "(i=>i) => o" where
+  mono_Ord :: "(i\<Rightarrow>i) \<Rightarrow> o" where
     "mono_Ord(F) \<equiv> \<forall>i j. i<j \<longrightarrow> F(i) < F(j)"
 
 definition
-  cont_Ord :: "(i=>i) => o" where
+  cont_Ord :: "(i\<Rightarrow>i) \<Rightarrow> o" where
     "cont_Ord(F) \<equiv> \<forall>l. Limit(l) \<longrightarrow> F(l) = (\<Union>i<l. F(i))"
 
 definition
-  Normal :: "(i=>i) => o" where
+  Normal :: "(i\<Rightarrow>i) \<Rightarrow> o" where
     "Normal(F) \<equiv> mono_Ord(F) \<and> cont_Ord(F)"
 
 
 subsubsection\<open>Immediate properties of the definitions\<close>
 
 lemma NormalI:
-     "\<lbrakk>\<And>i j. i<j \<Longrightarrow> F(i) < F(j);  \<And>l. Limit(l) \<Longrightarrow> F(l) = (\<Union>i<l. F(i))\<rbrakk>
+  "\<lbrakk>\<And>i j. i<j \<Longrightarrow> F(i) < F(j);  \<And>l. Limit(l) \<Longrightarrow> F(l) = (\<Union>i<l. F(i))\<rbrakk>
       \<Longrightarrow> Normal(F)"
-by (simp add: Normal_def mono_Ord_def cont_Ord_def)
+  by (simp add: Normal_def mono_Ord_def cont_Ord_def)
 
 lemma mono_Ord_imp_Ord: "\<lbrakk>Ord(i); mono_Ord(F)\<rbrakk> \<Longrightarrow> Ord(F(i))"
-apply (auto simp add: mono_Ord_def)
-apply (blast intro: lt_Ord) 
-done
+  apply (auto simp add: mono_Ord_def)
+  apply (blast intro: lt_Ord) 
+  done
 
 lemma mono_Ord_imp_mono: "\<lbrakk>i<j; mono_Ord(F)\<rbrakk> \<Longrightarrow> F(i) < F(j)"
-by (simp add: mono_Ord_def)
+  by (simp add: mono_Ord_def)
 
 lemma Normal_imp_Ord [simp]: "\<lbrakk>Normal(F); Ord(i)\<rbrakk> \<Longrightarrow> Ord(F(i))"
-by (simp add: Normal_def mono_Ord_imp_Ord) 
+  by (simp add: Normal_def mono_Ord_imp_Ord) 
 
 lemma Normal_imp_cont: "\<lbrakk>Normal(F); Limit(l)\<rbrakk> \<Longrightarrow> F(l) = (\<Union>i<l. F(i))"
-by (simp add: Normal_def cont_Ord_def)
+  by (simp add: Normal_def cont_Ord_def)
 
 lemma Normal_imp_mono: "\<lbrakk>i<j; Normal(F)\<rbrakk> \<Longrightarrow> F(i) < F(j)"
-by (simp add: Normal_def mono_Ord_def)
+  by (simp add: Normal_def mono_Ord_def)
 
 lemma Normal_increasing:
   assumes i: "Ord(i)" and F: "Normal(F)" shows"i \<le> F(i)"
@@ -272,58 +275,58 @@
 text\<open>The proof is from Drake, pages 113--114.\<close>
 
 lemma mono_Ord_imp_le_subset: "mono_Ord(F) \<Longrightarrow> mono_le_subset(F)"
-apply (simp add: mono_le_subset_def, clarify)
-apply (subgoal_tac "F(i)\<le>F(j)", blast dest: le_imp_subset) 
-apply (simp add: le_iff) 
-apply (blast intro: lt_Ord2 mono_Ord_imp_Ord mono_Ord_imp_mono) 
-done
+  apply (simp add: mono_le_subset_def, clarify)
+  apply (subgoal_tac "F(i)\<le>F(j)", blast dest: le_imp_subset) 
+  apply (simp add: le_iff) 
+  apply (blast intro: lt_Ord2 mono_Ord_imp_Ord mono_Ord_imp_mono) 
+  done
 
 text\<open>The following equation is taken for granted in any set theory text.\<close>
 lemma cont_Ord_Union:
-     "\<lbrakk>cont_Ord(F); mono_le_subset(F); X=0 \<longrightarrow> F(0)=0; \<forall>x\<in>X. Ord(x)\<rbrakk> 
+  "\<lbrakk>cont_Ord(F); mono_le_subset(F); X=0 \<longrightarrow> F(0)=0; \<forall>x\<in>X. Ord(x)\<rbrakk> 
       \<Longrightarrow> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
-apply (frule Ord_set_cases)
-apply (erule disjE, force) 
-apply (thin_tac "X=0 \<longrightarrow> Q" for Q, auto)
- txt\<open>The trival case of \<^term>\<open>\<Union>X \<in> X\<close>\<close>
- apply (rule equalityI, blast intro: Ord_Union_eq_succD) 
- apply (simp add: mono_le_subset_def UN_subset_iff le_subset_iff) 
- apply (blast elim: equalityE)
-txt\<open>The limit case, \<^term>\<open>Limit(\<Union>X)\<close>:
+  apply (frule Ord_set_cases)
+  apply (erule disjE, force) 
+  apply (thin_tac "X=0 \<longrightarrow> Q" for Q, auto)
+  txt\<open>The trival case of \<^term>\<open>\<Union>X \<in> X\<close>\<close>
+   apply (rule equalityI, blast intro: Ord_Union_eq_succD) 
+   apply (simp add: mono_le_subset_def UN_subset_iff le_subset_iff) 
+   apply (blast elim: equalityE)
+  txt\<open>The limit case, \<^term>\<open>Limit(\<Union>X)\<close>:
 @{subgoals[display,indent=0,margin=65]}
 \<close>
-apply (simp add: OUN_Union_eq cont_Ord_def)
-apply (rule equalityI) 
-txt\<open>First inclusion:\<close>
- apply (rule UN_least [OF OUN_least])
- apply (simp add: mono_le_subset_def, blast intro: leI) 
-txt\<open>Second inclusion:\<close>
-apply (rule UN_least) 
-apply (frule Union_upper_le, blast, blast)
-apply (erule leE, drule ltD, elim UnionE)
- apply (simp add: OUnion_def)
- apply blast+
-done
+  apply (simp add: OUN_Union_eq cont_Ord_def)
+  apply (rule equalityI) 
+  txt\<open>First inclusion:\<close>
+   apply (rule UN_least [OF OUN_least])
+   apply (simp add: mono_le_subset_def, blast intro: leI) 
+  txt\<open>Second inclusion:\<close>
+  apply (rule UN_least) 
+  apply (frule Union_upper_le, blast, blast)
+  apply (erule leE, drule ltD, elim UnionE)
+   apply (simp add: OUnion_def)
+   apply blast+
+  done
 
 lemma Normal_Union:
-     "\<lbrakk>X\<noteq>0; \<forall>x\<in>X. Ord(x); Normal(F)\<rbrakk> \<Longrightarrow> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
-apply (simp add: Normal_def) 
-apply (blast intro: mono_Ord_imp_le_subset cont_Ord_Union) 
-done
+  "\<lbrakk>X\<noteq>0; \<forall>x\<in>X. Ord(x); Normal(F)\<rbrakk> \<Longrightarrow> F(\<Union>(X)) = (\<Union>y\<in>X. F(y))"
+  apply (simp add: Normal_def) 
+  apply (blast intro: mono_Ord_imp_le_subset cont_Ord_Union) 
+  done
 
 lemma Normal_imp_fp_Closed: "Normal(F) \<Longrightarrow> Closed(\<lambda>i. F(i) = i)"
-apply (simp add: Closed_def ball_conj_distrib, clarify)
-apply (frule Ord_set_cases)
-apply (auto simp add: Normal_Union)
-done
+  apply (simp add: Closed_def ball_conj_distrib, clarify)
+  apply (frule Ord_set_cases)
+  apply (auto simp add: Normal_Union)
+  done
 
 
 lemma iterates_Normal_increasing:
-     "\<lbrakk>n\<in>nat;  x < F(x);  Normal(F)\<rbrakk> 
+  "\<lbrakk>n\<in>nat;  x < F(x);  Normal(F)\<rbrakk> 
       \<Longrightarrow> F^n (x) < F^(succ(n)) (x)"  
-apply (induct n rule: nat_induct)
-apply (simp_all add: Normal_imp_mono)
-done
+  apply (induct n rule: nat_induct)
+   apply (simp_all add: Normal_imp_mono)
+  done
 
 lemma Ord_iterates_Normal:
      "\<lbrakk>n\<in>nat;  Normal(F);  Ord(x)\<rbrakk> \<Longrightarrow> Ord(F^n (x))"  
@@ -331,20 +334,20 @@
 
 text\<open>THIS RESULT IS UNUSED\<close>
 lemma iterates_omega_Limit:
-     "\<lbrakk>Normal(F);  x < F(x)\<rbrakk> \<Longrightarrow> Limit(F^\<omega> (x))"  
-apply (frule lt_Ord) 
-apply (simp add: iterates_omega_def)
-apply (rule increasing_LimitI) 
-   \<comment> \<open>this lemma is @{thm increasing_LimitI [no_vars]}\<close>
- apply (blast intro: UN_upper_lt [of "1"]   Normal_imp_Ord
-                     Ord_iterates lt_imp_0_lt
-                     iterates_Normal_increasing, clarify)
-apply (rule bexI) 
- apply (blast intro: Ord_in_Ord [OF Ord_iterates_Normal]) 
-apply (rule UN_I, erule nat_succI) 
-apply (blast intro:  iterates_Normal_increasing Ord_iterates_Normal
-                     ltD [OF lt_trans1, OF succ_leI, OF ltI]) 
-done
+  "\<lbrakk>Normal(F);  x < F(x)\<rbrakk> \<Longrightarrow> Limit(F^\<omega> (x))"  
+  apply (frule lt_Ord) 
+  apply (simp add: iterates_omega_def)
+  apply (rule increasing_LimitI) 
+    \<comment> \<open>this lemma is @{thm increasing_LimitI [no_vars]}\<close>
+   apply (blast intro: UN_upper_lt [of "1"]   Normal_imp_Ord
+      Ord_iterates lt_imp_0_lt
+      iterates_Normal_increasing, clarify)
+  apply (rule bexI) 
+   apply (blast intro: Ord_in_Ord [OF Ord_iterates_Normal]) 
+  apply (rule UN_I, erule nat_succI) 
+  apply (blast intro:  iterates_Normal_increasing Ord_iterates_Normal
+      ltD [OF lt_trans1, OF succ_leI, OF ltI]) 
+  done
 
 lemma iterates_omega_fixedpoint:
      "\<lbrakk>Normal(F); Ord(a)\<rbrakk> \<Longrightarrow> F(F^\<omega> (a)) = F^\<omega> (a)" 
@@ -390,7 +393,7 @@
       normal function, by @{thm [source] Normal_imp_fp_Unbounded}.
 \<close>
 definition
-  normalize :: "[i=>i, i] => i" where
+  normalize :: "[i\<Rightarrow>i, i] \<Rightarrow> i" where
     "normalize(F,a) \<equiv> transrec2(a, F(0), \<lambda>x r. F(succ(x)) \<union> succ(r))"
 
 
@@ -428,11 +431,11 @@
 qed
 
 theorem Normal_normalize:
-     "(\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))) \<Longrightarrow> Normal(normalize(F))"
-apply (rule NormalI) 
-apply (blast intro!: normalize_increasing)
-apply (simp add: def_transrec2 [OF normalize_def])
-done
+  "(\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))) \<Longrightarrow> Normal(normalize(F))"
+  apply (rule NormalI) 
+   apply (blast intro!: normalize_increasing)
+  apply (simp add: def_transrec2 [OF normalize_def])
+  done
 
 theorem le_normalize:
   assumes a: "Ord(a)" and coF: "cont_Ord(F)" and F: "\<And>x. Ord(x) \<Longrightarrow> Ord(F(x))"
@@ -456,15 +459,15 @@
 numbers.\<close>
 
 definition
-  Aleph :: "i => i"  (\<open>\<aleph>_\<close> [90] 90) where
+  Aleph :: "i \<Rightarrow> i"  (\<open>\<aleph>_\<close> [90] 90) where
     "Aleph(a) \<equiv> transrec2(a, nat, \<lambda>x r. csucc(r))"
 
 lemma Card_Aleph [simp, intro]:
-     "Ord(a) \<Longrightarrow> Card(Aleph(a))"
-apply (erule trans_induct3) 
-apply (simp_all add: Card_csucc Card_nat Card_is_Ord
-                     def_transrec2 [OF Aleph_def])
-done
+  "Ord(a) \<Longrightarrow> Card(Aleph(a))"
+  apply (erule trans_induct3) 
+    apply (simp_all add: Card_csucc Card_nat Card_is_Ord
+      def_transrec2 [OF Aleph_def])
+  done
 
 lemma Aleph_increasing:
   assumes ab: "a < b" shows "Aleph(a) < Aleph(b)"
@@ -492,9 +495,9 @@
 qed
 
 theorem Normal_Aleph: "Normal(Aleph)"
-apply (rule NormalI) 
-apply (blast intro!: Aleph_increasing)
-apply (simp add: def_transrec2 [OF Aleph_def])
-done
+  apply (rule NormalI) 
+   apply (blast intro!: Aleph_increasing)
+  apply (simp add: def_transrec2 [OF Aleph_def])
+  done
 
 end
--- a/src/ZF/Constructible/Rank.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Rank.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -65,7 +65,7 @@
 
 lemma (in M_ordertype) wellordered_iso_pred_eq_lemma:
      "\<lbrakk>f \<in> \<langle>Order.pred(A,y,r), r\<rangle> \<cong> \<langle>Order.pred(A,x,r), r\<rangle>;
-       wellordered(M,A,r); x\<in>A; y\<in>A; M(A); M(f); M(r)\<rbrakk> \<Longrightarrow> <x,y> \<notin> r"
+       wellordered(M,A,r); x\<in>A; y\<in>A; M(A); M(f); M(r)\<rbrakk> \<Longrightarrow> \<langle>x,y\<rangle> \<notin> r"
 apply (frule wellordered_is_trans_on, assumption)
 apply (rule notI) 
 apply (drule_tac x2=y and x=x and r2=r in 
@@ -136,21 +136,21 @@
 
 
 definition  
-  obase :: "[i=>o,i,i] => i" where
+  obase :: "[i\<Rightarrow>o,i,i] \<Rightarrow> i" where
        \<comment> \<open>the domain of \<open>om\<close>, eventually shown to equal \<open>A\<close>\<close>
    "obase(M,A,r) \<equiv> {a\<in>A. \<exists>x[M]. \<exists>g[M]. Ord(x) \<and> 
                           g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x))}"
 
 definition
-  omap :: "[i=>o,i,i,i] => o" where
+  omap :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     \<comment> \<open>the function that maps wosets to order types\<close>
    "omap(M,A,r,f) \<equiv> 
         \<forall>z[M].
-         z \<in> f \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> \<and> Ord(x) \<and> 
+         z \<in> f \<longleftrightarrow> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = \<langle>a,x\<rangle> \<and> Ord(x) \<and> 
                         g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
 
 definition
-  otype :: "[i=>o,i,i,i] => o" where \<comment> \<open>the order types themselves\<close>
+  otype :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where \<comment> \<open>the order types themselves\<close>
    "otype(M,A,r,i) \<equiv> \<exists>f[M]. omap(M,A,r,f) \<and> is_range(M,f,i)"
 
 
@@ -160,7 +160,7 @@
 lemma (in M_ordertype) omap_iff:
      "\<lbrakk>omap(M,A,r,f); M(A); M(f)\<rbrakk> 
       \<Longrightarrow> z \<in> f \<longleftrightarrow>
-          (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> \<and> Ord(x) \<and> 
+          (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = \<langle>a,x\<rangle> \<and> Ord(x) \<and> 
                                 g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
 apply (simp add: omap_def) 
 apply (rule iffI)
@@ -416,16 +416,16 @@
    replacement below.  Here j is used to define the relation, namely
    Memrel(succ(j)), while x determines the domain of f.*)
 definition
-  is_oadd_fun :: "[i=>o,i,i,i,i] => o" where
+  is_oadd_fun :: "[i\<Rightarrow>o,i,i,i,i] \<Rightarrow> o" where
     "is_oadd_fun(M,i,j,x,f) \<equiv> 
        (\<forall>sj msj. M(sj) \<longrightarrow> M(msj) \<longrightarrow> 
                  successor(M,j,sj) \<longrightarrow> membership(M,sj,msj) \<longrightarrow> 
                  M_is_recfun(M, 
-                     %x g y. \<exists>gx[M]. image(M,g,x,gx) \<and> union(M,i,gx,y),
+                     \<lambda>x g y. \<exists>gx[M]. image(M,g,x,gx) \<and> union(M,i,gx,y),
                      msj, x, f))"
 
 definition
-  is_oadd :: "[i=>o,i,i,i] => o" where
+  is_oadd :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_oadd(M,i,j,k) \<equiv> 
         (\<not> ordinal(M,i) \<and> \<not> ordinal(M,j) \<and> k=0) |
         (\<not> ordinal(M,i) \<and> ordinal(M,j) \<and> k=j) |
@@ -437,7 +437,7 @@
 
 definition
  (*NEEDS RELATIVIZATION*)
-  omult_eqns :: "[i,i,i,i] => o" where
+  omult_eqns :: "[i,i,i,i] \<Rightarrow> o" where
     "omult_eqns(i,x,g,z) \<equiv>
             Ord(x) \<and> 
             (x=0 \<longrightarrow> z=0) \<and>
@@ -445,14 +445,14 @@
             (Limit(x) \<longrightarrow> z = \<Union>(g``x))"
 
 definition
-  is_omult_fun :: "[i=>o,i,i,i] => o" where
+  is_omult_fun :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_omult_fun(M,i,j,f) \<equiv> 
             (\<exists>df. M(df) \<and> is_function(M,f) \<and> 
                   is_domain(M,f,df) \<and> subset(M, j, df)) \<and> 
             (\<forall>x\<in>j. omult_eqns(i,x,f,f`x))"
 
 definition
-  is_omult :: "[i=>o,i,i,i] => o" where
+  is_omult :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_omult(M,i,j,k) \<equiv> 
         \<exists>f fj sj. M(f) \<and> M(fj) \<and> M(sj) \<and> 
                   successor(M,j,sj) \<and> is_omult_fun(M,i,sj,f) \<and> 
@@ -470,8 +470,8 @@
  and omult_strong_replacement':
    "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow>
     strong_replacement(M, 
-         \<lambda>x z. \<exists>y[M]. z = <x,y> \<and>
-             (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) \<and> 
+         \<lambda>x z. \<exists>y[M]. z = \<langle>x,y\<rangle> \<and>
+             (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,\<lambda>x g. THE z. omult_eqns(i,x,g,z),g) \<and> 
              y = (THE z. omult_eqns(i, x, g, z))))" 
 
 
@@ -483,7 +483,7 @@
         f \<in> a \<rightarrow> range(f) \<and> (\<forall>x. M(x) \<longrightarrow> x < a \<longrightarrow> f`x = i \<union> f``x)"
 apply (frule lt_Ord) 
 apply (simp add: is_oadd_fun_def  
-             relation2_def is_recfun_abs [of "%x g. i \<union> g``x"]
+             relation2_def is_recfun_abs [of "\<lambda>x g. i \<union> g``x"]
              is_recfun_iff_equation  
              Ball_def lt_trans [OF ltI, of _ a] lt_Memrel)
 apply (simp add: lt_def) 
@@ -494,18 +494,18 @@
 lemma (in M_ord_arith) oadd_strong_replacement':
     "\<lbrakk>M(i); M(j)\<rbrakk> \<Longrightarrow>
      strong_replacement(M, 
-            \<lambda>x z. \<exists>y[M]. z = <x,y> \<and>
-                  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i \<union> g``x,g) \<and> 
+            \<lambda>x z. \<exists>y[M]. z = \<langle>x,y\<rangle> \<and>
+                  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,\<lambda>x g. i \<union> g``x,g) \<and> 
                   y = i \<union> g``x))" 
 apply (insert oadd_strong_replacement [of i j]) 
 apply (simp add: is_oadd_fun_def relation2_def
-                 is_recfun_abs [of "%x g. i \<union> g``x"])  
+                 is_recfun_abs [of "\<lambda>x g. i \<union> g``x"])  
 done
 
 
 lemma (in M_ord_arith) exists_oadd:
     "\<lbrakk>Ord(j);  M(i);  M(j)\<rbrakk>
-     \<Longrightarrow> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. i \<union> g``x, f)"
+     \<Longrightarrow> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, \<lambda>x g. i \<union> g``x, f)"
 apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel])
     apply (simp_all add: Memrel_type oadd_strong_replacement') 
 done 
@@ -609,7 +609,7 @@
 
 lemma (in M_ord_arith) exists_omult:
     "\<lbrakk>Ord(j);  M(i);  M(j)\<rbrakk>
-     \<Longrightarrow> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. THE z. omult_eqns(i,x,g,z), f)"
+     \<Longrightarrow> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, \<lambda>x g. THE z. omult_eqns(i,x,g,z), f)"
 apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel])
     apply (simp_all add: Memrel_type omult_strong_replacement') 
 apply (blast intro: the_omult_eqns_closed) 
@@ -678,13 +678,13 @@
      "M(r) \<Longrightarrow>
       separation (M, \<lambda>x. 
          \<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow>
-         \<not> (\<exists>f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))"
+         \<not> (\<exists>f[M]. M_is_recfun(M, \<lambda>x f y. is_range(M,f,y), rplus, x, f)))"
  and wfrank_strong_replacement:
      "M(r) \<Longrightarrow>
       strong_replacement(M, \<lambda>x z. 
          \<forall>rplus[M]. tran_closure(M,r,rplus) \<longrightarrow>
          (\<exists>y[M]. \<exists>f[M]. pair(M,x,y,z)  \<and> 
-                        M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f) \<and>
+                        M_is_recfun(M, \<lambda>x f y. is_range(M,f,y), rplus, x, f) \<and>
                         is_range(M,f,y)))"
  and Ord_wfrank_separation:
      "M(r) \<Longrightarrow>
@@ -702,18 +702,18 @@
 lemma (in M_wfrank) wfrank_separation':
      "M(r) \<Longrightarrow>
       separation
-           (M, \<lambda>x. \<not> (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))"
+           (M, \<lambda>x. \<not> (\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. range(f), f)))"
 apply (insert wfrank_separation [of r])
-apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
+apply (simp add: relation2_def is_recfun_abs [of "\<lambda>x. range"])
 done
 
 lemma (in M_wfrank) wfrank_strong_replacement':
      "M(r) \<Longrightarrow>
       strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>f[M]. 
-                  pair(M,x,y,z) \<and> is_recfun(r^+, x, %x f. range(f), f) \<and>
+                  pair(M,x,y,z) \<and> is_recfun(r^+, x, \<lambda>x f. range(f), f) \<and>
                   y = range(f))"
 apply (insert wfrank_strong_replacement [of r])
-apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
+apply (simp add: relation2_def is_recfun_abs [of "\<lambda>x. range"])
 done
 
 lemma (in M_wfrank) Ord_wfrank_separation':
@@ -721,21 +721,21 @@
       separation (M, \<lambda>x. 
          \<not> (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) \<longrightarrow> Ord(range(f))))" 
 apply (insert Ord_wfrank_separation [of r])
-apply (simp add: relation2_def is_recfun_abs [of "%x. range"])
+apply (simp add: relation2_def is_recfun_abs [of "\<lambda>x. range"])
 done
 
 text\<open>This function, defined using replacement, is a rank function for
 well-founded relations within the class M.\<close>
 definition
-  wellfoundedrank :: "[i=>o,i,i] => i" where
+  wellfoundedrank :: "[i\<Rightarrow>o,i,i] \<Rightarrow> i" where
     "wellfoundedrank(M,r,A) \<equiv>
         {p. x\<in>A, \<exists>y[M]. \<exists>f[M]. 
-                       p = <x,y> \<and> is_recfun(r^+, x, %x f. range(f), f) \<and>
+                       p = \<langle>x,y\<rangle> \<and> is_recfun(r^+, x, \<lambda>x f. range(f), f) \<and>
                        y = range(f)}"
 
 lemma (in M_wfrank) exists_wfrank:
     "\<lbrakk>wellfounded(M,r); M(a); M(r)\<rbrakk>
-     \<Longrightarrow> \<exists>f[M]. is_recfun(r^+, a, %x f. range(f), f)"
+     \<Longrightarrow> \<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. range(f), f)"
 apply (rule wellfounded_exists_is_recfun)
       apply (blast intro: wellfounded_trancl)
      apply (rule trans_trancl)
@@ -759,7 +759,7 @@
 
 lemma (in M_wfrank) Ord_wfrank_range [rule_format]:
     "\<lbrakk>wellfounded(M,r); a\<in>A; M(r); M(A)\<rbrakk>
-     \<Longrightarrow> \<forall>f[M]. is_recfun(r^+, a, %x f. range(f), f) \<longrightarrow> Ord(range(f))"
+     \<Longrightarrow> \<forall>f[M]. is_recfun(r^+, a, \<lambda>x f. range(f), f) \<longrightarrow> Ord(range(f))"
 apply (drule wellfounded_trancl, assumption)
 apply (rule wellfounded_induct, assumption, erule (1) transM)
   apply simp
@@ -863,7 +863,7 @@
                  Ord_in_Ord [OF Ord_range_wellfoundedrank])
 
 lemma (in M_wfrank) wellfoundedrank_eq:
-     "\<lbrakk>is_recfun(r^+, a, %x. range, f);
+     "\<lbrakk>is_recfun(r^+, a, \<lambda>x. range, f);
          wellfounded(M,r);  a \<in> A; M(f); M(r); M(A)\<rbrakk>
       \<Longrightarrow> wellfoundedrank(M,r,A) ` a = range(f)"
 apply (rule apply_equality)
@@ -882,7 +882,7 @@
 
 
 lemma (in M_wfrank) wellfoundedrank_lt:
-     "\<lbrakk><a,b> \<in> r;
+     "\<lbrakk>\<langle>a,b\<rangle> \<in> r;
          wellfounded(M,r); r \<subseteq> A*A;  M(r); M(A)\<rbrakk>
       \<Longrightarrow> wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,A) ` b"
 apply (frule wellfounded_trancl, assumption)
--- a/src/ZF/Constructible/Rank_Separation.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Rank_Separation.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -133,17 +133,17 @@
 
 lemma wfrank_Reflects:
  "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
-              \<not> (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)),
+              \<not> (\<exists>f[L]. M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f)),
       \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
          \<not> (\<exists>f \<in> Lset(i).
-            M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y),
+            M_is_recfun(##Lset(i), \<lambda>x f y. is_range(##Lset(i),f,y),
                         rplus, x, f))]"
 by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection)
 
 lemma wfrank_separation:
      "L(r) \<Longrightarrow>
       separation (L, \<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
-         \<not> (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))"
+         \<not> (\<exists>f[L]. M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f)))"
 apply (rule gen_separation [OF wfrank_Reflects], simp)
 apply (rule_tac env="[r]" in DPow_LsetI)
 apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+
@@ -156,12 +156,12 @@
  "REFLECTS[\<lambda>z. \<exists>x[L]. x \<in> A \<and>
         (\<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
          (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z)  \<and>
-                        M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) \<and>
+                        M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<and>
                         is_range(L,f,y))),
  \<lambda>i z. \<exists>x \<in> Lset(i). x \<in> A \<and>
       (\<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
        (\<exists>y \<in> Lset(i). \<exists>f \<in> Lset(i). pair(##Lset(i),x,y,z)  \<and>
-         M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f) \<and>
+         M_is_recfun(##Lset(i), \<lambda>x f y. is_range(##Lset(i),f,y), rplus, x, f) \<and>
          is_range(##Lset(i),f,y)))]"
 by (intro FOL_reflections function_reflections fun_plus_reflections
              is_recfun_reflection tran_closure_reflection)
@@ -171,7 +171,7 @@
       strong_replacement(L, \<lambda>x z.
          \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
          (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z)  \<and>
-                        M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) \<and>
+                        M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<and>
                         is_range(L,f,y)))"
 apply (rule strong_replacementI)
 apply (rule_tac u="{r,B}" 
--- a/src/ZF/Constructible/Rec_Separation.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Rec_Separation.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -29,7 +29,7 @@
             fun_apply(M,f,j,fj) \<and> successor(M,j,sj) \<and>
             fun_apply(M,f,sj,fsj) \<and> pair(M,fj,fsj,ffp) \<and> ffp \<in> r)))"*)
 definition
-  rtran_closure_mem_fm :: "[i,i,i]=>i" where
+  rtran_closure_mem_fm :: "[i,i,i]\<Rightarrow>i" where
  "rtran_closure_mem_fm(A,r,p) \<equiv>
    Exists(Exists(Exists(
     And(omega_fm(2),
@@ -87,7 +87,7 @@
         \<forall>A[M]. is_field(M,r,A) \<longrightarrow>
          (\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))" *)
 definition
-  rtran_closure_fm :: "[i,i]=>i" where
+  rtran_closure_fm :: "[i,i]\<Rightarrow>i" where
   "rtran_closure_fm(r,s) \<equiv>
    Forall(Implies(field_fm(succ(r),0),
                   Forall(Iff(Member(0,succ(succ(s))),
@@ -122,7 +122,7 @@
 (*  "tran_closure(M,r,t) \<equiv>
          \<exists>s[M]. rtran_closure(M,r,s) \<and> composition(M,r,s,t)" *)
 definition
-  tran_closure_fm :: "[i,i]=>i" where
+  tran_closure_fm :: "[i,i]\<Rightarrow>i" where
   "tran_closure_fm(r,s) \<equiv>
    Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))"
 
@@ -295,7 +295,7 @@
 (* "is_nth(M,n,l,Z) \<equiv>
       \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) \<and> is_hd(M,X,Z)" *)
 definition
-  nth_fm :: "[i,i,i]=>i" where
+  nth_fm :: "[i,i,i]\<Rightarrow>i" where
     "nth_fm(n,l,Z) \<equiv> 
        Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0), 
               hd_fm(0,succ(Z))))"
--- a/src/ZF/Constructible/Reflection.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Reflection.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -24,7 +24,7 @@
 closed under ordered pairing provided \<open>l\<close> is limit.  Possibly this
 could be avoided: the induction hypothesis \<^term>\<open>Cl_reflects\<close>
 (in locale \<open>ex_reflection\<close>) could be weakened to
-\<^term>\<open>\<forall>y\<in>Mset(a). \<forall>z\<in>Mset(a). P(<y,z>) \<longleftrightarrow> Q(a,<y,z>)\<close>, removing most
+\<^term>\<open>\<forall>y\<in>Mset(a). \<forall>z\<in>Mset(a). P(\<langle>y,z\<rangle>) \<longleftrightarrow> Q(a,\<langle>y,z\<rangle>)\<close>, removing most
 uses of \<^term>\<open>Pair_in_Mset\<close>.  But there isn't much point in doing so, since
 ultimately the \<open>ex_reflection\<close> proof is packaged up using the
 predicate \<open>Reflects\<close>.
@@ -34,15 +34,15 @@
   assumes Mset_mono_le : "mono_le_subset(Mset)"
       and Mset_cont    : "cont_Ord(Mset)"
       and Pair_in_Mset : "\<lbrakk>x \<in> Mset(a); y \<in> Mset(a); Limit(a)\<rbrakk>
-                          \<Longrightarrow> <x,y> \<in> Mset(a)"
+                          \<Longrightarrow> \<langle>x,y\<rangle> \<in> Mset(a)"
   defines "M(x) \<equiv> \<exists>a. Ord(a) \<and> x \<in> Mset(a)"
       and "Reflects(Cl,P,Q) \<equiv> Closed_Unbounded(Cl) \<and>
                               (\<forall>a. Cl(a) \<longrightarrow> (\<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)))"
   fixes F0 \<comment> \<open>ordinal for a specific value \<^term>\<open>y\<close>\<close>
   fixes FF \<comment> \<open>sup over the whole level, \<^term>\<open>y\<in>Mset(a)\<close>\<close>
   fixes ClEx \<comment> \<open>Reflecting ordinals for the formula \<^term>\<open>\<exists>z. P\<close>\<close>
-  defines "F0(P,y) \<equiv> \<mu> b. (\<exists>z. M(z) \<and> P(<y,z>)) \<longrightarrow>
-                               (\<exists>z\<in>Mset(b). P(<y,z>))"
+  defines "F0(P,y) \<equiv> \<mu> b. (\<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)) \<longrightarrow>
+                               (\<exists>z\<in>Mset(b). P(\<langle>y,z\<rangle>))"
       and "FF(P)   \<equiv> \<lambda>a. \<Union>y\<in>Mset(a). F0(P,y)"
       and "ClEx(P,a) \<equiv> Limit(a) \<and> normalize(FF(P),a) = a"
 
@@ -97,7 +97,7 @@
 subsection\<open>Reflection for Existential Quantifiers\<close>
 
 lemma F0_works:
-  "\<lbrakk>y\<in>Mset(a); Ord(a); M(z); P(<y,z>)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(F0(P,y)). P(<y,z>)"
+  "\<lbrakk>y\<in>Mset(a); Ord(a); M(z); P(\<langle>y,z\<rangle>)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(F0(P,y)). P(\<langle>y,z\<rangle>)"
   unfolding F0_def M_def
   apply clarify
   apply (rule LeastI2)
@@ -117,7 +117,7 @@
 text\<open>Recall that \<^term>\<open>F0\<close> depends upon \<^term>\<open>y\<in>Mset(a)\<close>,
 while \<^term>\<open>FF\<close> depends only upon \<^term>\<open>a\<close>.\<close>
 lemma FF_works:
-  "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(FF(P,a)). P(<y,z>)"
+  "\<lbrakk>M(z); y\<in>Mset(a); P(\<langle>y,z\<rangle>); Ord(a)\<rbrakk> \<Longrightarrow> \<exists>z\<in>Mset(FF(P,a)). P(\<langle>y,z\<rangle>)"
   apply (simp add: FF_def)
   apply (simp_all add: cont_Ord_Union [of concl: Mset]
       Mset_cont Mset_mono_le not_emptyI)
@@ -125,8 +125,8 @@
   done
 
 lemma FFN_works:
-     "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Ord(a)\<rbrakk>
-      \<Longrightarrow> \<exists>z\<in>Mset(normalize(FF(P),a)). P(<y,z>)"
+     "\<lbrakk>M(z); y\<in>Mset(a); P(\<langle>y,z\<rangle>); Ord(a)\<rbrakk>
+      \<Longrightarrow> \<exists>z\<in>Mset(normalize(FF(P),a)). P(\<langle>y,z\<rangle>)"
 apply (drule FF_works [of concl: P], assumption+)
 apply (blast intro: cont_Ord_FF le_normalize [THEN Mset_mono, THEN subsetD])
 done
@@ -144,8 +144,8 @@
 begin
 
 lemma ClEx_downward:
-     "\<lbrakk>M(z); y\<in>Mset(a); P(<y,z>); Cl(a); ClEx(P,a)\<rbrakk>
-      \<Longrightarrow> \<exists>z\<in>Mset(a). Q(a,<y,z>)"
+     "\<lbrakk>M(z); y\<in>Mset(a); P(\<langle>y,z\<rangle>); Cl(a); ClEx(P,a)\<rbrakk>
+      \<Longrightarrow> \<exists>z\<in>Mset(a). Q(a,\<langle>y,z\<rangle>)"
 apply (simp add: ClEx_def, clarify)
 apply (frule Limit_is_Ord)
 apply (frule FFN_works [of concl: P], assumption+)
@@ -154,8 +154,8 @@
 done
 
 lemma ClEx_upward:
-     "\<lbrakk>z\<in>Mset(a); y\<in>Mset(a); Q(a,<y,z>); Cl(a); ClEx(P,a)\<rbrakk>
-      \<Longrightarrow> \<exists>z. M(z) \<and> P(<y,z>)"
+     "\<lbrakk>z\<in>Mset(a); y\<in>Mset(a); Q(a,\<langle>y,z\<rangle>); Cl(a); ClEx(P,a)\<rbrakk>
+      \<Longrightarrow> \<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)"
 apply (simp add: ClEx_def M_def)
 apply (blast dest: Cl_reflects
              intro: Limit_is_Ord Pair_in_Mset)
@@ -164,7 +164,7 @@
 text\<open>Class \<open>ClEx\<close> indeed consists of reflecting ordinals...\<close>
 lemma ZF_ClEx_iff:
      "\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a)\<rbrakk>
-      \<Longrightarrow> (\<exists>z. M(z) \<and> P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+      \<Longrightarrow> (\<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,\<langle>y,z\<rangle>))"
 by (blast intro: dest: ClEx_downward ClEx_upward)
 
 text\<open>...and it is closed and unbounded\<close>
@@ -186,7 +186,7 @@
 lemma ClEx_iff:
      "\<lbrakk>y\<in>Mset(a); Cl(a); ClEx(P,a);
         \<And>a. \<lbrakk>Cl(a); Ord(a)\<rbrakk> \<Longrightarrow> \<forall>x\<in>Mset(a). P(x) \<longleftrightarrow> Q(a,x)\<rbrakk>
-      \<Longrightarrow> (\<exists>z. M(z) \<and> P(<y,z>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,<y,z>))"
+      \<Longrightarrow> (\<exists>z. M(z) \<and> P(\<langle>y,z\<rangle>)) \<longleftrightarrow> (\<exists>z\<in>Mset(a). Q(a,\<langle>y,z\<rangle>))"
 apply (unfold ClEx_def FF_def F0_def M_def)
 apply (rule ex_reflection.ZF_ClEx_iff
   [OF ex_reflection.intro, OF reflection.intro ex_reflection_axioms.intro,
@@ -215,8 +215,8 @@
 lemma Ex_reflection_0:
      "Reflects(Cl,P0,Q0)
       \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(P0,a),
-                   \<lambda>x. \<exists>z. M(z) \<and> P0(<x,z>),
-                   \<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,<x,z>))"
+                   \<lambda>x. \<exists>z. M(z) \<and> P0(\<langle>x,z\<rangle>),
+                   \<lambda>a x. \<exists>z\<in>Mset(a). Q0(a,\<langle>x,z\<rangle>))"
 apply (simp add: Reflects_def)
 apply (intro conjI Closed_Unbounded_Int)
   apply blast
@@ -227,8 +227,8 @@
 lemma All_reflection_0:
      "Reflects(Cl,P0,Q0)
       \<Longrightarrow> Reflects(\<lambda>a. Cl(a) \<and> ClEx(\<lambda>x.\<not>P0(x), a),
-                   \<lambda>x. \<forall>z. M(z) \<longrightarrow> P0(<x,z>),
-                   \<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,<x,z>))"
+                   \<lambda>x. \<forall>z. M(z) \<longrightarrow> P0(\<langle>x,z\<rangle>),
+                   \<lambda>a x. \<forall>z\<in>Mset(a). Q0(a,\<langle>x,z\<rangle>))"
 apply (simp only: all_iff_not_ex_not ball_iff_not_bex_not)
 apply (rule Not_reflection, drule Not_reflection, simp)
 apply (erule Ex_reflection_0)
--- a/src/ZF/Constructible/Relative.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Relative.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -11,120 +11,120 @@
 subsection\<open>Relativized versions of standard set-theoretic concepts\<close>
 
 definition
-  empty :: "[i=>o,i] => o" where
+  empty :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "empty(M,z) \<equiv> \<forall>x[M]. x \<notin> z"
 
 definition
-  subset :: "[i=>o,i,i] => o" where
+  subset :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "subset(M,A,B) \<equiv> \<forall>x[M]. x\<in>A \<longrightarrow> x \<in> B"
 
 definition
-  upair :: "[i=>o,i,i,i] => o" where
+  upair :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "upair(M,a,b,z) \<equiv> a \<in> z \<and> b \<in> z \<and> (\<forall>x[M]. x\<in>z \<longrightarrow> x = a \<or> x = b)"
 
 definition
-  pair :: "[i=>o,i,i,i] => o" where
+  pair :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "pair(M,a,b,z) \<equiv> \<exists>x[M]. upair(M,a,a,x) \<and>
                      (\<exists>y[M]. upair(M,a,b,y) \<and> upair(M,x,y,z))"
 
 
 definition
-  union :: "[i=>o,i,i,i] => o" where
+  union :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "union(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a \<or> x \<in> b"
 
 definition
-  is_cons :: "[i=>o,i,i,i] => o" where
+  is_cons :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_cons(M,a,b,z) \<equiv> \<exists>x[M]. upair(M,a,a,x) \<and> union(M,x,b,z)"
 
 definition
-  successor :: "[i=>o,i,i] => o" where
+  successor :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "successor(M,a,z) \<equiv> is_cons(M,a,a,z)"
 
 definition
-  number1 :: "[i=>o,i] => o" where
+  number1 :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "number1(M,a) \<equiv> \<exists>x[M]. empty(M,x) \<and> successor(M,x,a)"
 
 definition
-  number2 :: "[i=>o,i] => o" where
+  number2 :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "number2(M,a) \<equiv> \<exists>x[M]. number1(M,x) \<and> successor(M,x,a)"
 
 definition
-  number3 :: "[i=>o,i] => o" where
+  number3 :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "number3(M,a) \<equiv> \<exists>x[M]. number2(M,x) \<and> successor(M,x,a)"
 
 definition
-  powerset :: "[i=>o,i,i] => o" where
+  powerset :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "powerset(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> subset(M,x,A)"
 
 definition
-  is_Collect :: "[i=>o,i,i=>o,i] => o" where
+  is_Collect :: "[i\<Rightarrow>o,i,i\<Rightarrow>o,i] \<Rightarrow> o" where
     "is_Collect(M,A,P,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A \<and> P(x)"
 
 definition
-  is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" where
+  is_Replace :: "[i\<Rightarrow>o,i,[i,i]\<Rightarrow>o,i] \<Rightarrow> o" where
     "is_Replace(M,A,P,z) \<equiv> \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A \<and> P(x,u))"
 
 definition
-  inter :: "[i=>o,i,i,i] => o" where
+  inter :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "inter(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a \<and> x \<in> b"
 
 definition
-  setdiff :: "[i=>o,i,i,i] => o" where
+  setdiff :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "setdiff(M,a,b,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a \<and> x \<notin> b"
 
 definition
-  big_union :: "[i=>o,i,i] => o" where
+  big_union :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "big_union(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A \<and> x \<in> y)"
 
 definition
-  big_inter :: "[i=>o,i,i] => o" where
+  big_inter :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "big_inter(M,A,z) \<equiv>
              (A=0 \<longrightarrow> z=0) \<and>
              (A\<noteq>0 \<longrightarrow> (\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> x \<in> y)))"
 
 definition
-  cartprod :: "[i=>o,i,i,i] => o" where
+  cartprod :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "cartprod(M,A,B,z) \<equiv>
         \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A \<and> (\<exists>y[M]. y\<in>B \<and> pair(M,x,y,u)))"
 
 definition
-  is_sum :: "[i=>o,i,i,i] => o" where
+  is_sum :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_sum(M,A,B,Z) \<equiv>
        \<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
        number1(M,n1) \<and> cartprod(M,n1,A,A0) \<and> upair(M,n1,n1,s1) \<and>
        cartprod(M,s1,B,B1) \<and> union(M,A0,B1,Z)"
 
 definition
-  is_Inl :: "[i=>o,i,i] => o" where
+  is_Inl :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_Inl(M,a,z) \<equiv> \<exists>zero[M]. empty(M,zero) \<and> pair(M,zero,a,z)"
 
 definition
-  is_Inr :: "[i=>o,i,i] => o" where
+  is_Inr :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_Inr(M,a,z) \<equiv> \<exists>n1[M]. number1(M,n1) \<and> pair(M,n1,a,z)"
 
 definition
-  is_converse :: "[i=>o,i,i] => o" where
+  is_converse :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_converse(M,r,z) \<equiv>
         \<forall>x[M]. x \<in> z \<longleftrightarrow>
              (\<exists>w[M]. w\<in>r \<and> (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) \<and> pair(M,v,u,x)))"
 
 definition
-  pre_image :: "[i=>o,i,i,i] => o" where
+  pre_image :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "pre_image(M,r,A,z) \<equiv>
         \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>y[M]. y\<in>A \<and> pair(M,x,y,w)))"
 
 definition
-  is_domain :: "[i=>o,i,i] => o" where
+  is_domain :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_domain(M,r,z) \<equiv>
         \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>y[M]. pair(M,x,y,w)))"
 
 definition
-  image :: "[i=>o,i,i,i] => o" where
+  image :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "image(M,r,A,z) \<equiv>
         \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>x[M]. x\<in>A \<and> pair(M,x,y,w)))"
 
 definition
-  is_range :: "[i=>o,i,i] => o" where
+  is_range :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     \<comment> \<open>the cleaner
       \<^term>\<open>\<exists>r'[M]. is_converse(M,r,r') \<and> is_domain(M,r',z)\<close>
       unfortunately needs an instance of separation in order to prove
@@ -133,41 +133,41 @@
         \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>x[M]. pair(M,x,y,w)))"
 
 definition
-  is_field :: "[i=>o,i,i] => o" where
+  is_field :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "is_field(M,r,z) \<equiv>
         \<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) \<and> is_range(M,r,rr) \<and>
                         union(M,dr,rr,z)"
 
 definition
-  is_relation :: "[i=>o,i] => o" where
+  is_relation :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "is_relation(M,r) \<equiv>
         (\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))"
 
 definition
-  is_function :: "[i=>o,i] => o" where
+  is_function :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "is_function(M,r) \<equiv>
         \<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
            pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'"
 
 definition
-  fun_apply :: "[i=>o,i,i,i] => o" where
+  fun_apply :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "fun_apply(M,f,x,y) \<equiv>
         (\<exists>xs[M]. \<exists>fxs[M].
          upair(M,x,x,xs) \<and> image(M,f,xs,fxs) \<and> big_union(M,fxs,y))"
 
 definition
-  typed_function :: "[i=>o,i,i,i] => o" where
+  typed_function :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "typed_function(M,A,B,r) \<equiv>
         is_function(M,r) \<and> is_relation(M,r) \<and> is_domain(M,r,A) \<and>
         (\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))"
 
 definition
-  is_funspace :: "[i=>o,i,i,i] => o" where
+  is_funspace :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "is_funspace(M,A,B,F) \<equiv>
         \<forall>f[M]. f \<in> F \<longleftrightarrow> typed_function(M,A,B,f)"
 
 definition
-  composition :: "[i=>o,i,i,i] => o" where
+  composition :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "composition(M,r,s,t) \<equiv>
         \<forall>p[M]. p \<in> t \<longleftrightarrow>
                (\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
@@ -175,104 +175,104 @@
                 xy \<in> s \<and> yz \<in> r)"
 
 definition
-  injection :: "[i=>o,i,i,i] => o" where
+  injection :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "injection(M,A,B,f) \<equiv>
         typed_function(M,A,B,f) \<and>
         (\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
           pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')"
 
 definition
-  surjection :: "[i=>o,i,i,i] => o" where
+  surjection :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "surjection(M,A,B,f) \<equiv>
         typed_function(M,A,B,f) \<and>
         (\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A \<and> fun_apply(M,f,x,y)))"
 
 definition
-  bijection :: "[i=>o,i,i,i] => o" where
+  bijection :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "bijection(M,A,B,f) \<equiv> injection(M,A,B,f) \<and> surjection(M,A,B,f)"
 
 definition
-  restriction :: "[i=>o,i,i,i] => o" where
+  restriction :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     "restriction(M,r,A,z) \<equiv>
         \<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r \<and> (\<exists>u[M]. u\<in>A \<and> (\<exists>v[M]. pair(M,u,v,x))))"
 
 definition
-  transitive_set :: "[i=>o,i] => o" where
+  transitive_set :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "transitive_set(M,a) \<equiv> \<forall>x[M]. x\<in>a \<longrightarrow> subset(M,x,a)"
 
 definition
-  ordinal :: "[i=>o,i] => o" where
+  ordinal :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
      \<comment> \<open>an ordinal is a transitive set of transitive sets\<close>
     "ordinal(M,a) \<equiv> transitive_set(M,a) \<and> (\<forall>x[M]. x\<in>a \<longrightarrow> transitive_set(M,x))"
 
 definition
-  limit_ordinal :: "[i=>o,i] => o" where
+  limit_ordinal :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     \<comment> \<open>a limit ordinal is a non-empty, successor-closed ordinal\<close>
     "limit_ordinal(M,a) \<equiv>
         ordinal(M,a) \<and> \<not> empty(M,a) \<and>
         (\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a \<and> successor(M,x,y)))"
 
 definition
-  successor_ordinal :: "[i=>o,i] => o" where
+  successor_ordinal :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     \<comment> \<open>a successor ordinal is any ordinal that is neither empty nor limit\<close>
     "successor_ordinal(M,a) \<equiv>
         ordinal(M,a) \<and> \<not> empty(M,a) \<and> \<not> limit_ordinal(M,a)"
 
 definition
-  finite_ordinal :: "[i=>o,i] => o" where
+  finite_ordinal :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     \<comment> \<open>an ordinal is finite if neither it nor any of its elements are limit\<close>
     "finite_ordinal(M,a) \<equiv>
         ordinal(M,a) \<and> \<not> limit_ordinal(M,a) \<and>
         (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))"
 
 definition
-  omega :: "[i=>o,i] => o" where
+  omega :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     \<comment> \<open>omega is a limit ordinal none of whose elements are limit\<close>
     "omega(M,a) \<equiv> limit_ordinal(M,a) \<and> (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))"
 
 definition
-  is_quasinat :: "[i=>o,i] => o" where
+  is_quasinat :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "is_quasinat(M,z) \<equiv> empty(M,z) \<or> (\<exists>m[M]. successor(M,m,z))"
 
 definition
-  is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" where
+  is_nat_case :: "[i\<Rightarrow>o, i, [i,i]\<Rightarrow>o, i, i] \<Rightarrow> o" where
     "is_nat_case(M, a, is_b, k, z) \<equiv>
        (empty(M,k) \<longrightarrow> z=a) \<and>
        (\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) \<and>
        (is_quasinat(M,k) \<or> empty(M,z))"
 
 definition
-  relation1 :: "[i=>o, [i,i]=>o, i=>i] => o" where
+  relation1 :: "[i\<Rightarrow>o, [i,i]\<Rightarrow>o, i\<Rightarrow>i] \<Rightarrow> o" where
     "relation1(M,is_f,f) \<equiv> \<forall>x[M]. \<forall>y[M]. is_f(x,y) \<longleftrightarrow> y = f(x)"
 
 definition
-  Relation1 :: "[i=>o, i, [i,i]=>o, i=>i] => o" where
+  Relation1 :: "[i\<Rightarrow>o, i, [i,i]\<Rightarrow>o, i\<Rightarrow>i] \<Rightarrow> o" where
     \<comment> \<open>as above, but typed\<close>
     "Relation1(M,A,is_f,f) \<equiv>
         \<forall>x[M]. \<forall>y[M]. x\<in>A \<longrightarrow> is_f(x,y) \<longleftrightarrow> y = f(x)"
 
 definition
-  relation2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o" where
+  relation2 :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, [i,i]\<Rightarrow>i] \<Rightarrow> o" where
     "relation2(M,is_f,f) \<equiv> \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
 
 definition
-  Relation2 :: "[i=>o, i, i, [i,i,i]=>o, [i,i]=>i] => o" where
+  Relation2 :: "[i\<Rightarrow>o, i, i, [i,i,i]\<Rightarrow>o, [i,i]\<Rightarrow>i] \<Rightarrow> o" where
     "Relation2(M,A,B,is_f,f) \<equiv>
         \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
 
 definition
-  relation3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
+  relation3 :: "[i\<Rightarrow>o, [i,i,i,i]\<Rightarrow>o, [i,i,i]\<Rightarrow>i] \<Rightarrow> o" where
     "relation3(M,is_f,f) \<equiv>
        \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M]. is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
 
 definition
-  Relation3 :: "[i=>o, i, i, i, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
+  Relation3 :: "[i\<Rightarrow>o, i, i, i, [i,i,i,i]\<Rightarrow>o, [i,i,i]\<Rightarrow>i] \<Rightarrow> o" where
     "Relation3(M,A,B,C,is_f,f) \<equiv>
        \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M].
          x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> z\<in>C \<longrightarrow> is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
 
 definition
-  relation4 :: "[i=>o, [i,i,i,i,i]=>o, [i,i,i,i]=>i] => o" where
+  relation4 :: "[i\<Rightarrow>o, [i,i,i,i,i]\<Rightarrow>o, [i,i,i,i]\<Rightarrow>i] \<Rightarrow> o" where
     "relation4(M,is_f,f) \<equiv>
        \<forall>u[M]. \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>a[M]. is_f(u,x,y,z,a) \<longleftrightarrow> a = f(u,x,y,z)"
 
@@ -290,12 +290,12 @@
 subsection \<open>The relativized ZF axioms\<close>
 
 definition
-  extensionality :: "(i=>o) => o" where
+  extensionality :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "extensionality(M) \<equiv>
         \<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x \<longleftrightarrow> z \<in> y) \<longrightarrow> x=y"
 
 definition
-  separation :: "[i=>o, i=>o] => o" where
+  separation :: "[i\<Rightarrow>o, i\<Rightarrow>o] \<Rightarrow> o" where
     \<comment> \<open>The formula \<open>P\<close> should only involve parameters
         belonging to \<open>M\<close> and all its quantifiers must be relativized
         to \<open>M\<close>.  We do not have separation as a scheme; every instance
@@ -304,36 +304,36 @@
         \<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z \<and> P(x)"
 
 definition
-  upair_ax :: "(i=>o) => o" where
+  upair_ax :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "upair_ax(M) \<equiv> \<forall>x[M]. \<forall>y[M]. \<exists>z[M]. upair(M,x,y,z)"
 
 definition
-  Union_ax :: "(i=>o) => o" where
+  Union_ax :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "Union_ax(M) \<equiv> \<forall>x[M]. \<exists>z[M]. big_union(M,x,z)"
 
 definition
-  power_ax :: "(i=>o) => o" where
+  power_ax :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "power_ax(M) \<equiv> \<forall>x[M]. \<exists>z[M]. powerset(M,x,z)"
 
 definition
-  univalent :: "[i=>o, i, [i,i]=>o] => o" where
+  univalent :: "[i\<Rightarrow>o, i, [i,i]\<Rightarrow>o] \<Rightarrow> o" where
     "univalent(M,A,P) \<equiv>
         \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. \<forall>z[M]. P(x,y) \<and> P(x,z) \<longrightarrow> y=z)"
 
 definition
-  replacement :: "[i=>o, [i,i]=>o] => o" where
+  replacement :: "[i\<Rightarrow>o, [i,i]\<Rightarrow>o] \<Rightarrow> o" where
     "replacement(M,P) \<equiv>
       \<forall>A[M]. univalent(M,A,P) \<longrightarrow>
       (\<exists>Y[M]. \<forall>b[M]. (\<exists>x[M]. x\<in>A \<and> P(x,b)) \<longrightarrow> b \<in> Y)"
 
 definition
-  strong_replacement :: "[i=>o, [i,i]=>o] => o" where
+  strong_replacement :: "[i\<Rightarrow>o, [i,i]\<Rightarrow>o] \<Rightarrow> o" where
     "strong_replacement(M,P) \<equiv>
       \<forall>A[M]. univalent(M,A,P) \<longrightarrow>
       (\<exists>Y[M]. \<forall>b[M]. b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A \<and> P(x,b)))"
 
 definition
-  foundation_ax :: "(i=>o) => o" where
+  foundation_ax :: "(i\<Rightarrow>o) \<Rightarrow> o" where
     "foundation_ax(M) \<equiv>
         \<forall>x[M]. (\<exists>y[M]. y\<in>x) \<longrightarrow> (\<exists>y[M]. y\<in>x \<and> \<not>(\<exists>z[M]. z\<in>x \<and> z \<in> y))"
 
@@ -360,12 +360,12 @@
 text\<open>Congruence rule for separation: can assume the variable is in \<open>M\<close>\<close>
 lemma separation_cong [cong]:
      "(\<And>x. M(x) \<Longrightarrow> P(x) \<longleftrightarrow> P'(x))
-      \<Longrightarrow> separation(M, %x. P(x)) \<longleftrightarrow> separation(M, %x. P'(x))"
+      \<Longrightarrow> separation(M, \<lambda>x. P(x)) \<longleftrightarrow> separation(M, \<lambda>x. P'(x))"
 by (simp add: separation_def)
 
 lemma univalent_cong [cong]:
      "\<lbrakk>A=A'; \<And>x y. \<lbrakk>x\<in>A; M(x); M(y)\<rbrakk> \<Longrightarrow> P(x,y) \<longleftrightarrow> P'(x,y)\<rbrakk>
-      \<Longrightarrow> univalent(M, A, %x y. P(x,y)) \<longleftrightarrow> univalent(M, A', %x y. P'(x,y))"
+      \<Longrightarrow> univalent(M, A, \<lambda>x y. P(x,y)) \<longleftrightarrow> univalent(M, A', \<lambda>x y. P'(x,y))"
 by (simp add: univalent_def)
 
 lemma univalent_triv [intro,simp]:
@@ -379,8 +379,8 @@
 text\<open>Congruence rule for replacement\<close>
 lemma strong_replacement_cong [cong]:
      "\<lbrakk>\<And>x y. \<lbrakk>M(x); M(y)\<rbrakk> \<Longrightarrow> P(x,y) \<longleftrightarrow> P'(x,y)\<rbrakk>
-      \<Longrightarrow> strong_replacement(M, %x y. P(x,y)) \<longleftrightarrow>
-          strong_replacement(M, %x y. P'(x,y))"
+      \<Longrightarrow> strong_replacement(M, \<lambda>x y. P(x,y)) \<longleftrightarrow>
+          strong_replacement(M, \<lambda>x y. P'(x,y))"
 by (simp add: strong_replacement_def)
 
 text\<open>The extensionality axiom\<close>
@@ -459,13 +459,13 @@
 subsection\<open>Lemmas Needed to Reduce Some Set Constructions to Instances
       of Separation\<close>
 
-lemma image_iff_Collect: "r `` A = {y \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>x\<in>A. p=<x,y>}"
+lemma image_iff_Collect: "r `` A = {y \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>x\<in>A. p=\<langle>x,y\<rangle>}"
 apply (rule equalityI, auto)
 apply (simp add: Pair_def, blast)
 done
 
 lemma vimage_iff_Collect:
-     "r -`` A = {x \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>y\<in>A. p=<x,y>}"
+     "r -`` A = {x \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>y\<in>A. p=\<langle>x,y\<rangle>}"
 apply (rule equalityI, auto)
 apply (simp add: Pair_def, blast)
 done
@@ -503,7 +503,7 @@
 text\<open>More constants, for order types\<close>
 
 definition
-  order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where
+  order_isomorphism :: "[i\<Rightarrow>o,i,i,i,i,i] \<Rightarrow> o" where
     "order_isomorphism(M,A,r,B,s,f) \<equiv>
         bijection(M,A,B,f) \<and>
         (\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
@@ -512,12 +512,12 @@
             pair(M,fx,fy,q) \<longrightarrow> (p\<in>r \<longleftrightarrow> q\<in>s))))"
 
 definition
-  pred_set :: "[i=>o,i,i,i,i] => o" where
+  pred_set :: "[i\<Rightarrow>o,i,i,i,i] \<Rightarrow> o" where
     "pred_set(M,A,x,r,B) \<equiv>
         \<forall>y[M]. y \<in> B \<longleftrightarrow> (\<exists>p[M]. p\<in>r \<and> y \<in> A \<and> pair(M,y,x,p))"
 
 definition
-  membership :: "[i=>o,i,i] => o" where \<comment> \<open>membership relation\<close>
+  membership :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where \<comment> \<open>membership relation\<close>
     "membership(M,A,r) \<equiv>
         \<forall>p[M]. p \<in> r \<longleftrightarrow> (\<exists>x[M]. x\<in>A \<and> (\<exists>y[M]. y\<in>A \<and> x\<in>y \<and> pair(M,x,y,p)))"
 
@@ -628,25 +628,25 @@
   by blast
 
 lemma (in M_trans) pair_abs [simp]:
-     "M(z) \<Longrightarrow> pair(M,a,b,z) \<longleftrightarrow> z=<a,b>"
+     "M(z) \<Longrightarrow> pair(M,a,b,z) \<longleftrightarrow> z=\<langle>a,b\<rangle>"
 apply (simp add: pair_def Pair_def)
 apply (blast intro: transM)
 done
 
 lemma (in M_trans) pair_in_MD [dest!]:
-     "M(<a,b>) \<Longrightarrow> M(a) \<and> M(b)"
+     "M(\<langle>a,b\<rangle>) \<Longrightarrow> M(a) \<and> M(b)"
   by (simp add: Pair_def, blast intro: transM)
 
 lemma (in M_trivial) pair_in_MI [intro!]:
-     "M(a) \<and> M(b) \<Longrightarrow> M(<a,b>)"
+     "M(a) \<and> M(b) \<Longrightarrow> M(\<langle>a,b\<rangle>)"
   by (simp add: Pair_def)
 
 lemma (in M_trivial) pair_in_M_iff [simp]:
-     "M(<a,b>) \<longleftrightarrow> M(a) \<and> M(b)"
+     "M(\<langle>a,b\<rangle>) \<longleftrightarrow> M(a) \<and> M(b)"
   by blast
 
 lemma (in M_trans) pair_components_in_M:
-     "\<lbrakk><x,y> \<in> A; M(A)\<rbrakk> \<Longrightarrow> M(x) \<and> M(y)"
+     "\<lbrakk>\<langle>x,y\<rangle> \<in> A; M(A)\<rbrakk> \<Longrightarrow> M(x) \<and> M(y)"
   by (blast dest: transM)
 
 lemma (in M_trivial) cartprod_abs [simp]:
@@ -739,8 +739,8 @@
      "\<lbrakk>A=A';
          \<And>x y. \<lbrakk>M(x); M(y)\<rbrakk> \<Longrightarrow> P(x,y) \<longleftrightarrow> P'(x,y);
          z=z'\<rbrakk>
-      \<Longrightarrow> is_Replace(M, A, %x y. P(x,y), z) \<longleftrightarrow>
-          is_Replace(M, A', %x y. P'(x,y), z')"
+      \<Longrightarrow> is_Replace(M, A, \<lambda>x y. P(x,y), z) \<longleftrightarrow>
+          is_Replace(M, A', \<lambda>x y. P'(x,y), z')"
 by (simp add: is_Replace_def)
 
 lemma (in M_trans) univalent_Replace_iff:
@@ -795,7 +795,7 @@
       makes relativization easier.\<close>
 lemma (in M_trans) RepFun_closed2:
      "\<lbrakk>strong_replacement(M, \<lambda>x y. x\<in>A \<and> y = f(x)); M(A); \<forall>x\<in>A. M(f(x))\<rbrakk>
-      \<Longrightarrow> M(RepFun(A, %x. f(x)))"
+      \<Longrightarrow> M(RepFun(A, \<lambda>x. f(x)))"
 apply (simp add: RepFun_def)
 apply (frule strong_replacement_closed, assumption)
 apply (auto dest: transM  simp add: Replace_conj_eq univalent_def)
@@ -804,7 +804,7 @@
 subsubsection \<open>Absoluteness for \<^term>\<open>Lambda\<close>\<close>
 
 definition
- is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where
+ is_lambda :: "[i\<Rightarrow>o, i, [i,i]\<Rightarrow>o, i] \<Rightarrow> o" where
     "is_lambda(M, A, is_b, z) \<equiv>
        \<forall>p[M]. p \<in> z \<longleftrightarrow>
         (\<exists>u[M]. \<exists>v[M]. u\<in>A \<and> pair(M,u,v,p) \<and> is_b(u,v))"
@@ -838,8 +838,8 @@
 lemma is_lambda_cong [cong]:
      "\<lbrakk>A=A';  z=z';
          \<And>x y. \<lbrakk>x\<in>A; M(x); M(y)\<rbrakk> \<Longrightarrow> is_b(x,y) \<longleftrightarrow> is_b'(x,y)\<rbrakk>
-      \<Longrightarrow> is_lambda(M, A, %x y. is_b(x,y), z) \<longleftrightarrow>
-          is_lambda(M, A', %x y. is_b'(x,y), z')"
+      \<Longrightarrow> is_lambda(M, A, \<lambda>x y. is_b(x,y), z) \<longleftrightarrow>
+          is_lambda(M, A', \<lambda>x y. is_b'(x,y), z')"
 by (simp add: is_lambda_def)
 
 lemma (in M_trans) image_abs [simp]:
@@ -1001,7 +1001,7 @@
   equations only hold for x\<in>nat (or in some other set) and not for the
   whole of the class M.
   consts
-    natnumber_aux :: "[i=>o,i] => i"
+    natnumber_aux :: "[i\<Rightarrow>o,i] \<Rightarrow> i"
 
   primrec
       "natnumber_aux(M,0) = (\<lambda>x\<in>nat. if empty(M,x) then 1 else 0)"
@@ -1010,7 +1010,7 @@
                      then 1 else 0)"
 
   definition
-    natnumber :: "[i=>o,i,i] => o"
+    natnumber :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o"
       "natnumber(M,n,x) \<equiv> natnumber_aux(M,n)`x = 1"
 
   lemma (in M_trivial) [simp]:
@@ -1076,7 +1076,7 @@
      "\<lbrakk>M(A); M(B); M(C)\<rbrakk>
       \<Longrightarrow> cartprod(M,A,B,C) \<longleftrightarrow>
           (\<exists>p1[M]. \<exists>p2[M]. powerset(M,A \<union> B,p1) \<and> powerset(M,p1,p2) \<and>
-                   C = {z \<in> p2. \<exists>x\<in>A. \<exists>y\<in>B. z = <x,y>})"
+                   C = {z \<in> p2. \<exists>x\<in>A. \<exists>y\<in>B. z = \<langle>x,y\<rangle>})"
 apply (simp add: Pair_def cartprod_def, safe)
 defer 1
   apply (simp add: powerset_def)
@@ -1348,15 +1348,15 @@
 
 (*???equalities*)
 lemma Collect_Un_Collect_eq:
-     "Collect(A,P) \<union> Collect(A,Q) = Collect(A, %x. P(x) \<or> Q(x))"
+     "Collect(A,P) \<union> Collect(A,Q) = Collect(A, \<lambda>x. P(x) \<or> Q(x))"
 by blast
 
 lemma Diff_Collect_eq:
-     "A - Collect(A,P) = Collect(A, %x. \<not> P(x))"
+     "A - Collect(A,P) = Collect(A, \<lambda>x. \<not> P(x))"
 by blast
 
 lemma (in M_trans) Collect_rall_eq:
-     "M(Y) \<Longrightarrow> Collect(A, %x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y)) =
+     "M(Y) \<Longrightarrow> Collect(A, \<lambda>x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y)) =
                (if Y=0 then A else (\<Inter>y \<in> Y. {x \<in> A. P(x,y)}))"
   by (simp,blast dest: transM)
 
@@ -1405,7 +1405,7 @@
 lemma (in M_basic) succ_fun_eq2:
      "\<lbrakk>M(B); M(n->B)\<rbrakk> \<Longrightarrow>
       succ(n) -> B =
-      \<Union>{z. p \<in> (n->B)*B, \<exists>f[M]. \<exists>b[M]. p = <f,b> \<and> z = {cons(<n,b>, f)}}"
+      \<Union>{z. p \<in> (n->B)*B, \<exists>f[M]. \<exists>b[M]. p = \<langle>f,b\<rangle> \<and> z = {cons(\<langle>n,b\<rangle>, f)}}"
 apply (simp add: succ_fun_eq)
 apply (blast dest: transM)
 done
@@ -1429,21 +1429,21 @@
 subsection\<open>Relativization and Absoluteness for Boolean Operators\<close>
 
 definition
-  is_bool_of_o :: "[i=>o, o, i] => o" where
+  is_bool_of_o :: "[i\<Rightarrow>o, o, i] \<Rightarrow> o" where
    "is_bool_of_o(M,P,z) \<equiv> (P \<and> number1(M,z)) \<or> (\<not>P \<and> empty(M,z))"
 
 definition
-  is_not :: "[i=>o, i, i] => o" where
+  is_not :: "[i\<Rightarrow>o, i, i] \<Rightarrow> o" where
    "is_not(M,a,z) \<equiv> (number1(M,a)  \<and> empty(M,z)) |
                      (\<not>number1(M,a) \<and> number1(M,z))"
 
 definition
-  is_and :: "[i=>o, i, i, i] => o" where
+  is_and :: "[i\<Rightarrow>o, i, i, i] \<Rightarrow> o" where
    "is_and(M,a,b,z) \<equiv> (number1(M,a)  \<and> z=b) |
                        (\<not>number1(M,a) \<and> empty(M,z))"
 
 definition
-  is_or :: "[i=>o, i, i, i] => o" where
+  is_or :: "[i\<Rightarrow>o, i, i, i] \<Rightarrow> o" where
    "is_or(M,a,b,z) \<equiv> (number1(M,a)  \<and> number1(M,z)) |
                       (\<not>number1(M,a) \<and> z=b)"
 
@@ -1485,12 +1485,12 @@
 subsection\<open>Relativization and Absoluteness for List Operators\<close>
 
 definition
-  is_Nil :: "[i=>o, i] => o" where
+  is_Nil :: "[i\<Rightarrow>o, i] \<Rightarrow> o" where
      \<comment> \<open>because \<^prop>\<open>[] \<equiv> Inl(0)\<close>\<close>
     "is_Nil(M,xs) \<equiv> \<exists>zero[M]. empty(M,zero) \<and> is_Inl(M,zero,xs)"
 
 definition
-  is_Cons :: "[i=>o,i,i,i] => o" where
+  is_Cons :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
      \<comment> \<open>because \<^prop>\<open>Cons(a, l) \<equiv> Inr(\<langle>a,l\<rangle>)\<close>\<close>
     "is_Cons(M,a,l,Z) \<equiv> \<exists>p[M]. pair(M,a,l,p) \<and> is_Inr(M,p,Z)"
 
@@ -1510,21 +1510,21 @@
 
 
 definition
-  quasilist :: "i => o" where
+  quasilist :: "i \<Rightarrow> o" where
     "quasilist(xs) \<equiv> xs=Nil \<or> (\<exists>x l. xs = Cons(x,l))"
 
 definition
-  is_quasilist :: "[i=>o,i] => o" where
+  is_quasilist :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where
     "is_quasilist(M,z) \<equiv> is_Nil(M,z) \<or> (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))"
 
 definition
-  list_case' :: "[i, [i,i]=>i, i] => i" where
+  list_case' :: "[i, [i,i]\<Rightarrow>i, i] \<Rightarrow> i" where
     \<comment> \<open>A version of \<^term>\<open>list_case\<close> that's always defined.\<close>
     "list_case'(a,b,xs) \<equiv>
        if quasilist(xs) then list_case(a,b,xs) else 0"
 
 definition
-  is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o" where
+  is_list_case :: "[i\<Rightarrow>o, i, [i,i,i]\<Rightarrow>o, i, i] \<Rightarrow> o" where
     \<comment> \<open>Returns 0 for non-lists\<close>
     "is_list_case(M, a, is_b, xs, z) \<equiv>
        (is_Nil(M,xs) \<longrightarrow> z=a) \<and>
@@ -1532,17 +1532,17 @@
        (is_quasilist(M,xs) \<or> empty(M,z))"
 
 definition
-  hd' :: "i => i" where
+  hd' :: "i \<Rightarrow> i" where
     \<comment> \<open>A version of \<^term>\<open>hd\<close> that's always defined.\<close>
     "hd'(xs) \<equiv> if quasilist(xs) then hd(xs) else 0"
 
 definition
-  tl' :: "i => i" where
+  tl' :: "i \<Rightarrow> i" where
     \<comment> \<open>A version of \<^term>\<open>tl\<close> that's always defined.\<close>
     "tl'(xs) \<equiv> if quasilist(xs) then tl(xs) else 0"
 
 definition
-  is_hd :: "[i=>o,i,i] => o" where
+  is_hd :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
      \<comment> \<open>\<^term>\<open>hd([]) = 0\<close> no constraints if not a list.
           Avoiding implication prevents the simplifier's looping.\<close>
     "is_hd(M,xs,H) \<equiv>
@@ -1551,7 +1551,7 @@
        (is_quasilist(M,xs) \<or> empty(M,H))"
 
 definition
-  is_tl :: "[i=>o,i,i] => o" where
+  is_tl :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
      \<comment> \<open>\<^term>\<open>tl([]) = []\<close>; see comments about \<^term>\<open>is_hd\<close>\<close>
     "is_tl(M,xs,T) \<equiv>
        (is_Nil(M,xs) \<longrightarrow> T=xs) \<and>
--- a/src/ZF/Constructible/Satisfies_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Satisfies_absolute.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -17,7 +17,7 @@
         is_formula_N(M,n,formula_n) \<and> p \<notin> formula_n \<and>
         successor(M,n,sn) \<and> is_formula_N(M,sn,formula_sn) \<and> p \<in> formula_sn" *)
 definition
-  depth_fm :: "[i,i]=>i" where
+  depth_fm :: "[i,i]\<Rightarrow>i" where
   "depth_fm(p,n) \<equiv> 
      Exists(Exists(Exists(
        And(formula_N_fm(n#+3,1),
@@ -58,7 +58,7 @@
       will be enclosed by three quantifiers.\<close>
 
 (* is_formula_case :: 
-    "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o"
+    "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, [i,i,i]\<Rightarrow>o, [i,i,i]\<Rightarrow>o, [i,i]\<Rightarrow>o, i, i] \<Rightarrow> o"
   "is_formula_case(M, is_a, is_b, is_c, is_d, v, z) \<equiv> 
       (\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> is_Member(M,x,y,v) \<longrightarrow> is_a(x,y,z)) \<and>
       (\<forall>x[M]. \<forall>y[M]. x\<in>nat \<longrightarrow> y\<in>nat \<longrightarrow> is_Equal(M,x,y,v) \<longrightarrow> is_b(x,y,z)) \<and>
@@ -67,7 +67,7 @@
       (\<forall>x[M]. x\<in>formula \<longrightarrow> is_Forall(M,x,v) \<longrightarrow> is_d(x,z))" *)
 
 definition
-  formula_case_fm :: "[i, i, i, i, i, i]=>i" where
+  formula_case_fm :: "[i, i, i, i, i, i]\<Rightarrow>i" where
   "formula_case_fm(is_a, is_b, is_c, is_d, v, z) \<equiv> 
         And(Forall(Forall(Implies(finite_ordinal_fm(1), 
                            Implies(finite_ordinal_fm(0), 
@@ -175,7 +175,7 @@
 subsection \<open>Absoluteness for the Function \<^term>\<open>satisfies\<close>\<close>
 
 definition
-  is_depth_apply :: "[i=>o,i,i,i] => o" where
+  is_depth_apply :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
    \<comment> \<open>Merely a useful abbreviation for the sequel.\<close>
   "is_depth_apply(M,h,p,z) \<equiv>
     \<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M]. 
@@ -195,12 +195,12 @@
 text\<open>These constants let us instantiate the parameters \<^term>\<open>a\<close>, \<^term>\<open>b\<close>,
       \<^term>\<open>c\<close>, \<^term>\<open>d\<close>, etc., of the locale \<open>Formula_Rec\<close>.\<close>
 definition
-  satisfies_a :: "[i,i,i]=>i" where
+  satisfies_a :: "[i,i,i]\<Rightarrow>i" where
    "satisfies_a(A) \<equiv> 
     \<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) \<in> nth(y,env))"
 
 definition
-  satisfies_is_a :: "[i=>o,i,i,i,i]=>o" where
+  satisfies_is_a :: "[i\<Rightarrow>o,i,i,i,i]\<Rightarrow>o" where
    "satisfies_is_a(M,A) \<equiv> 
     \<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
@@ -210,12 +210,12 @@
                 zz)"
 
 definition
-  satisfies_b :: "[i,i,i]=>i" where
+  satisfies_b :: "[i,i,i]\<Rightarrow>i" where
    "satisfies_b(A) \<equiv>
     \<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) = nth(y,env))"
 
 definition
-  satisfies_is_b :: "[i=>o,i,i,i,i]=>o" where
+  satisfies_is_b :: "[i\<Rightarrow>o,i,i,i,i]\<Rightarrow>o" where
    \<comment> \<open>We simplify the formula to have just \<^term>\<open>nx\<close> rather than 
        introducing \<^term>\<open>ny\<close> with  \<^term>\<open>nx=ny\<close>\<close>
   "satisfies_is_b(M,A) \<equiv> 
@@ -226,11 +226,11 @@
                 zz)"
 
 definition 
-  satisfies_c :: "[i,i,i,i,i]=>i" where
+  satisfies_c :: "[i,i,i,i,i]\<Rightarrow>i" where
    "satisfies_c(A) \<equiv> \<lambda>p q rp rq. \<lambda>env \<in> list(A). not(rp ` env and rq ` env)"
 
 definition
-  satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o" where
+  satisfies_is_c :: "[i\<Rightarrow>o,i,i,i,i,i]\<Rightarrow>o" where
    "satisfies_is_c(M,A,h) \<equiv> 
     \<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M]. 
@@ -240,12 +240,12 @@
                 zz)"
 
 definition
-  satisfies_d :: "[i,i,i]=>i" where
+  satisfies_d :: "[i,i,i]\<Rightarrow>i" where
    "satisfies_d(A) 
     \<equiv> \<lambda>p rp. \<lambda>env \<in> list(A). bool_of_o (\<forall>x\<in>A. rp ` (Cons(x,env)) = 1)"
 
 definition
-  satisfies_is_d :: "[i=>o,i,i,i,i]=>o" where
+  satisfies_is_d :: "[i\<Rightarrow>o,i,i,i,i]\<Rightarrow>o" where
    "satisfies_is_d(M,A,h) \<equiv> 
     \<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) \<longrightarrow>
              is_lambda(M, lA, 
@@ -258,7 +258,7 @@
                zz)"
 
 definition
-  satisfies_MH :: "[i=>o,i,i,i,i]=>o" where
+  satisfies_MH :: "[i\<Rightarrow>o,i,i,i,i]\<Rightarrow>o" where
     \<comment> \<open>The variable \<^term>\<open>u\<close> is unused, but gives \<^term>\<open>satisfies_MH\<close> 
         the correct arity.\<close>
   "satisfies_MH \<equiv> 
@@ -271,7 +271,7 @@
                z)"
 
 definition
-  is_satisfies :: "[i=>o,i,i,i]=>o" where
+  is_satisfies :: "[i\<Rightarrow>o,i,i,i]\<Rightarrow>o" where
   "is_satisfies(M,A) \<equiv> is_formula_rec (M, satisfies_MH(M,A))"
 
 
@@ -515,7 +515,7 @@
         finite_ordinal(M,dp) \<and> is_depth(M,p,dp) \<and> successor(M,dp,sdp) \<and>
         fun_apply(M,h,sdp,hsdp) \<and> fun_apply(M,hsdp,p,z) *)
 definition
-  depth_apply_fm :: "[i,i,i]=>i" where
+  depth_apply_fm :: "[i,i,i]\<Rightarrow>i" where
     "depth_apply_fm(h,p,z) \<equiv>
        Exists(Exists(Exists(
         And(finite_ordinal_fm(2),
@@ -559,7 +559,7 @@
                 zz)  *)
 
 definition
-  satisfies_is_a_fm :: "[i,i,i,i]=>i" where
+  satisfies_is_a_fm :: "[i,i,i,i]\<Rightarrow>i" where
   "satisfies_is_a_fm(A,x,y,z) \<equiv>
    Forall(
      Implies(is_list_fm(succ(A),0),
@@ -611,7 +611,7 @@
                 zz) *)
 
 definition
-  satisfies_is_b_fm :: "[i,i,i,i]=>i" where
+  satisfies_is_b_fm :: "[i,i,i,i]\<Rightarrow>i" where
  "satisfies_is_b_fm(A,x,y,z) \<equiv>
    Forall(
      Implies(is_list_fm(succ(A),0),
@@ -661,7 +661,7 @@
                 zz) *)
 
 definition
-  satisfies_is_c_fm :: "[i,i,i,i,i]=>i" where
+  satisfies_is_c_fm :: "[i,i,i,i,i]\<Rightarrow>i" where
  "satisfies_is_c_fm(A,h,p,q,zz) \<equiv>
    Forall(
      Implies(is_list_fm(succ(A),0),
@@ -715,7 +715,7 @@
                zz) *)
 
 definition
-  satisfies_is_d_fm :: "[i,i,i,i]=>i" where
+  satisfies_is_d_fm :: "[i,i,i,i]\<Rightarrow>i" where
  "satisfies_is_d_fm(A,h,p,zz) \<equiv>
    Forall(
      Implies(is_list_fm(succ(A),0),
@@ -770,7 +770,7 @@
                zz) *)
 
 definition
-  satisfies_MH_fm :: "[i,i,i,i]=>i" where
+  satisfies_MH_fm :: "[i,i,i,i]\<Rightarrow>i" where
  "satisfies_MH_fm(A,u,f,zz) \<equiv>
    Forall(
      Implies(is_formula_fm(0),
--- a/src/ZF/Constructible/WF_absolute.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/WF_absolute.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,10 +9,10 @@
 subsection\<open>Transitive closure without fixedpoints\<close>
 
 definition
-  rtrancl_alt :: "[i,i]=>i" where
+  rtrancl_alt :: "[i,i]\<Rightarrow>i" where
     "rtrancl_alt(A,r) \<equiv>
        {p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A.
-                 (\<exists>x y. p = <x,y> \<and>  f`0 = x \<and> f`n = y) \<and>
+                 (\<exists>x y. p = \<langle>x,y\<rangle> \<and>  f`0 = x \<and> f`n = y) \<and>
                        (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}"
 
 lemma alt_rtrancl_lemma1 [rule_format]:
@@ -59,7 +59,7 @@
 
 
 definition
-  rtran_closure_mem :: "[i=>o,i,i,i] => o" where
+  rtran_closure_mem :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" where
     \<comment> \<open>The property of belonging to \<open>rtran_closure(r)\<close>\<close>
     "rtran_closure_mem(M,A,r,p) \<equiv>
               \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
@@ -73,13 +73,13 @@
                       fun_apply(M,f,sj,fsj) \<and> pair(M,fj,fsj,ffp) \<and> ffp \<in> r)))"
 
 definition
-  rtran_closure :: "[i=>o,i,i] => o" where
+  rtran_closure :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "rtran_closure(M,r,s) \<equiv> 
         \<forall>A[M]. is_field(M,r,A) \<longrightarrow>
          (\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))"
 
 definition
-  tran_closure :: "[i=>o,i,i] => o" where
+  tran_closure :: "[i\<Rightarrow>o,i,i] \<Rightarrow> o" where
     "tran_closure(M,r,t) \<equiv>
          \<exists>s[M]. rtran_closure(M,r,s) \<and> composition(M,r,s,t)"
     
@@ -98,7 +98,7 @@
       \<Longrightarrow> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
           (\<exists>n[M]. n\<in>nat \<and> 
            (\<exists>f[M]. f \<in> succ(n) -> A \<and>
-            (\<exists>x[M]. \<exists>y[M]. p = <x,y> \<and> f`0 = x \<and> f`n = y) \<and>
+            (\<exists>x[M]. \<exists>y[M]. p = \<langle>x,y\<rangle> \<and> f`0 = x \<and> f`n = y) \<and>
                            (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)))"
   apply (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) 
 done
@@ -137,7 +137,7 @@
 by (simp add: tran_closure_def trancl_def)
 
 lemma (in M_trancl) wellfounded_trancl_separation':
-     "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z \<and> <w,x> \<in> r^+)"
+     "\<lbrakk>M(r); M(Z)\<rbrakk> \<Longrightarrow> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z \<and> \<langle>w,x\<rangle> \<in> r^+)"
 by (insert wellfounded_trancl_separation [of r Z], simp) 
 
 text\<open>Alternative proof of \<open>wf_on_trancl\<close>; inspiration for the
--- a/src/ZF/Constructible/WFrec.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/WFrec.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,7 +12,7 @@
 (*Many of these might be useful in WF.thy*)
 
 lemma apply_recfun2:
-    "\<lbrakk>is_recfun(r,a,H,f); <x,i>:f\<rbrakk> \<Longrightarrow> i = H(x, restrict(f,r-``{x}))"
+    "\<lbrakk>is_recfun(r,a,H,f); \<langle>x,i\<rangle>:f\<rbrakk> \<Longrightarrow> i = H(x, restrict(f,r-``{x}))"
 apply (frule apply_recfun) 
  apply (blast dest: is_recfun_type fun_is_rel) 
 apply (simp add: function_apply_equality [OF _ is_recfun_imp_function])
@@ -36,7 +36,7 @@
 by (blast dest: is_recfun_type fun_is_rel)
 
 lemma trans_Int_eq:
-      "\<lbrakk>trans(r); <y,x> \<in> r\<rbrakk> \<Longrightarrow> r -`` {x} \<inter> r -`` {y} = r -`` {y}"
+      "\<lbrakk>trans(r); \<langle>y,x\<rangle> \<in> r\<rbrakk> \<Longrightarrow> r -`` {x} \<inter> r -`` {y} = r -`` {y}"
 by (blast intro: transD) 
 
 lemma is_recfun_restrict_idem:
@@ -91,7 +91,7 @@
     "\<lbrakk>is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  
        wellfounded(M,r);  trans(r);
        M(f); M(g); M(r); M(x); M(a); M(b)\<rbrakk> 
-     \<Longrightarrow> <x,a> \<in> r \<longrightarrow> <x,b> \<in> r \<longrightarrow> f`x=g`x"
+     \<Longrightarrow> \<langle>x,a\<rangle> \<in> r \<longrightarrow> \<langle>x,b\<rangle> \<in> r \<longrightarrow> f`x=g`x"
 apply (frule_tac f=f in is_recfun_type) 
 apply (frule_tac f=g in is_recfun_type) 
 apply (simp add: is_recfun_def)
@@ -103,8 +103,8 @@
 apply (erule ssubst)+
 apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def)
 apply (rename_tac x1)
-apply (rule_tac t="%z. H(x1,z)" in subst_context) 
-apply (subgoal_tac "\<forall>y \<in> r-``{x1}. \<forall>z. <y,z>\<in>f \<longleftrightarrow> <y,z>\<in>g")
+apply (rule_tac t="\<lambda>z. H(x1,z)" in subst_context) 
+apply (subgoal_tac "\<forall>y \<in> r-``{x1}. \<forall>z. \<langle>y,z\<rangle>\<in>f \<longleftrightarrow> \<langle>y,z\<rangle>\<in>g")
  apply (blast intro: transD) 
 apply (simp add: apply_iff) 
 apply (blast intro: transD sym) 
@@ -113,7 +113,7 @@
 lemma (in M_basic) is_recfun_cut: 
     "\<lbrakk>is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  
        wellfounded(M,r); trans(r); 
-       M(f); M(g); M(r); <b,a> \<in> r\<rbrakk>   
+       M(f); M(g); M(r); \<langle>b,a\<rangle> \<in> r\<rbrakk>   
       \<Longrightarrow> restrict(f, r-``{b}) = g"
 apply (frule_tac f=f in is_recfun_type) 
 apply (rule fun_extension) 
@@ -135,7 +135,7 @@
   "\<lbrakk>M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
    \<Longrightarrow> is_recfun(r,a,H,f) \<longleftrightarrow>
        (\<forall>z[M]. z \<in> f \<longleftrightarrow> 
-        (\<exists>x[M]. <x,a> \<in> r \<and> z = <x, H(x, restrict(f, r-``{x}))>))"
+        (\<exists>x[M]. \<langle>x,a\<rangle> \<in> r \<and> z = <x, H(x, restrict(f, r-``{x}))>))"
 apply (simp add: is_recfun_def lam_def)
 apply (safe intro!: equalityI) 
    apply (drule equalityD1 [THEN subsetD], assumption) 
@@ -174,21 +174,21 @@
        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));  M(Y);
        \<forall>b[M]. 
            b \<in> Y \<longleftrightarrow>
-           (\<exists>x[M]. <x,a1> \<in> r \<and>
+           (\<exists>x[M]. \<langle>x,a1\<rangle> \<in> r \<and>
             (\<exists>y[M]. b = \<langle>x,y\<rangle> \<and> (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))));
           \<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f)\<rbrakk>
        \<Longrightarrow> restrict(Y, r -`` {x}) = f"
-apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. <y,z>:Y \<longleftrightarrow> <y,z>:f") 
+apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. \<langle>y,z\<rangle>:Y \<longleftrightarrow> \<langle>y,z\<rangle>:f") 
  apply (simp (no_asm_simp) add: restrict_def) 
  apply (thin_tac "rall(M,P)" for P)+  \<comment> \<open>essential for efficiency\<close>
  apply (frule is_recfun_type [THEN fun_is_rel], blast)
 apply (frule pair_components_in_M, assumption, clarify) 
 apply (rule iffI)
- apply (frule_tac y="<y,z>" in transM, assumption)
+ apply (frule_tac y="\<langle>y,z\<rangle>" in transM, assumption)
  apply (clarsimp simp add: vimage_singleton_iff is_recfun_type [THEN apply_iff]
                            apply_recfun is_recfun_cut) 
 txt\<open>Opposite inclusion: something in f, show in Y\<close>
-apply (frule_tac y="<y,z>" in transM, assumption)  
+apply (frule_tac y="\<langle>y,z\<rangle>" in transM, assumption)  
 apply (simp add: vimage_singleton_iff) 
 apply (rule conjI) 
  apply (blast dest: transD) 
@@ -271,7 +271,7 @@
 subsection\<open>Relativization of the ZF Predicate \<^term>\<open>is_recfun\<close>\<close>
 
 definition
-  M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where
+  M_is_recfun :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i, i] \<Rightarrow> o" where
    "M_is_recfun(M,MH,r,a,f) \<equiv> 
      \<forall>z[M]. z \<in> f \<longleftrightarrow> 
             (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. 
@@ -280,12 +280,12 @@
                xa \<in> r \<and> MH(x, f_r_sx, y))"
 
 definition
-  is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where
+  is_wfrec :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i, i] \<Rightarrow> o" where
    "is_wfrec(M,MH,r,a,z) \<equiv> 
       \<exists>f[M]. M_is_recfun(M,MH,r,a,f) \<and> MH(a,f,z)"
 
 definition
-  wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where
+  wfrec_replacement :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i] \<Rightarrow> o" where
    "wfrec_replacement(M,MH,r) \<equiv>
         strong_replacement(M, 
              \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) \<and> is_wfrec(M,MH,r,x,y))"
@@ -324,8 +324,8 @@
 lemma wfrec_replacement_cong [cong]:
      "\<lbrakk>\<And>x y z. \<lbrakk>M(x); M(y); M(z)\<rbrakk> \<Longrightarrow> MH(x,y,z) \<longleftrightarrow> MH'(x,y,z);
          r=r'\<rbrakk> 
-      \<Longrightarrow> wfrec_replacement(M, %x y. MH(x,y), r) \<longleftrightarrow> 
-          wfrec_replacement(M, %x y. MH'(x,y), r')" 
+      \<Longrightarrow> wfrec_replacement(M, \<lambda>x y. MH(x,y), r) \<longleftrightarrow> 
+          wfrec_replacement(M, \<lambda>x y. MH'(x,y), r')" 
 by (simp add: is_wfrec_def wfrec_replacement_def) 
 
 
--- a/src/ZF/Constructible/Wellorderings.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Constructible/Wellorderings.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -16,33 +16,33 @@
 subsection\<open>Wellorderings\<close>
 
 definition
-  irreflexive :: "[i=>o,i,i]=>o" where
-    "irreflexive(M,A,r) \<equiv> \<forall>x[M]. x\<in>A \<longrightarrow> <x,x> \<notin> r"
+  irreflexive :: "[i\<Rightarrow>o,i,i]\<Rightarrow>o" where
+    "irreflexive(M,A,r) \<equiv> \<forall>x[M]. x\<in>A \<longrightarrow> \<langle>x,x\<rangle> \<notin> r"
   
 definition
-  transitive_rel :: "[i=>o,i,i]=>o" where
+  transitive_rel :: "[i\<Rightarrow>o,i,i]\<Rightarrow>o" where
     "transitive_rel(M,A,r) \<equiv> 
         \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> (\<forall>z[M]. z\<in>A \<longrightarrow> 
-                          <x,y>\<in>r \<longrightarrow> <y,z>\<in>r \<longrightarrow> <x,z>\<in>r))"
+                          \<langle>x,y\<rangle>\<in>r \<longrightarrow> \<langle>y,z\<rangle>\<in>r \<longrightarrow> \<langle>x,z\<rangle>\<in>r))"
 
 definition
-  linear_rel :: "[i=>o,i,i]=>o" where
+  linear_rel :: "[i\<Rightarrow>o,i,i]\<Rightarrow>o" where
     "linear_rel(M,A,r) \<equiv> 
-        \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> <x,y>\<in>r | x=y | <y,x>\<in>r)"
+        \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> \<langle>x,y\<rangle>\<in>r | x=y | \<langle>y,x\<rangle>\<in>r)"
 
 definition
-  wellfounded :: "[i=>o,i]=>o" where
+  wellfounded :: "[i\<Rightarrow>o,i]\<Rightarrow>o" where
     \<comment> \<open>EVERY non-empty set has an \<open>r\<close>-minimal element\<close>
     "wellfounded(M,r) \<equiv> 
-        \<forall>x[M]. x\<noteq>0 \<longrightarrow> (\<exists>y[M]. y\<in>x \<and> \<not>(\<exists>z[M]. z\<in>x \<and> <z,y> \<in> r))"
+        \<forall>x[M]. x\<noteq>0 \<longrightarrow> (\<exists>y[M]. y\<in>x \<and> \<not>(\<exists>z[M]. z\<in>x \<and> \<langle>z,y\<rangle> \<in> r))"
 definition
-  wellfounded_on :: "[i=>o,i,i]=>o" where
+  wellfounded_on :: "[i\<Rightarrow>o,i,i]\<Rightarrow>o" where
     \<comment> \<open>every non-empty SUBSET OF \<open>A\<close> has an \<open>r\<close>-minimal element\<close>
     "wellfounded_on(M,A,r) \<equiv> 
-        \<forall>x[M]. x\<noteq>0 \<longrightarrow> x\<subseteq>A \<longrightarrow> (\<exists>y[M]. y\<in>x \<and> \<not>(\<exists>z[M]. z\<in>x \<and> <z,y> \<in> r))"
+        \<forall>x[M]. x\<noteq>0 \<longrightarrow> x\<subseteq>A \<longrightarrow> (\<exists>y[M]. y\<in>x \<and> \<not>(\<exists>z[M]. z\<in>x \<and> \<langle>z,y\<rangle> \<in> r))"
 
 definition
-  wellordered :: "[i=>o,i,i]=>o" where
+  wellordered :: "[i\<Rightarrow>o,i,i]\<Rightarrow>o" where
     \<comment> \<open>linear and wellfounded on \<open>A\<close>\<close>
     "wellordered(M,A,r) \<equiv> 
         transitive_rel(M,A,r) \<and> linear_rel(M,A,r) \<and> wellfounded_on(M,A,r)"
@@ -108,7 +108,7 @@
 (*Consider the least z in domain(r) such that P(z) does not hold...*)
 lemma (in M_basic) wellfounded_induct: 
      "\<lbrakk>wellfounded(M,r); M(a); M(r); separation(M, \<lambda>x. \<not>P(x));  
-         \<forall>x. M(x) \<and> (\<forall>y. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x)\<rbrakk>
+         \<forall>x. M(x) \<and> (\<forall>y. \<langle>y,x\<rangle> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x)\<rbrakk>
       \<Longrightarrow> P(a)"
 apply (simp (no_asm_use) add: wellfounded_def)
 apply (drule_tac x="{z \<in> domain(r). \<not>P(z)}" in rspec)
@@ -118,7 +118,7 @@
 lemma (in M_basic) wellfounded_on_induct: 
      "\<lbrakk>a\<in>A;  wellfounded_on(M,A,r);  M(A);  
        separation(M, \<lambda>x. x\<in>A \<longrightarrow> \<not>P(x));  
-       \<forall>x\<in>A. M(x) \<and> (\<forall>y\<in>A. <y,x> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x)\<rbrakk>
+       \<forall>x\<in>A. M(x) \<and> (\<forall>y\<in>A. \<langle>y,x\<rangle> \<in> r \<longrightarrow> P(y)) \<longrightarrow> P(x)\<rbrakk>
       \<Longrightarrow> P(a)"
 apply (simp (no_asm_use) add: wellfounded_on_def)
 apply (drule_tac x="{z\<in>A. z\<in>A \<longrightarrow> \<not>P(z)}" in rspec)
@@ -182,7 +182,7 @@
    apply (frule transM, assumption)
    apply blast
   apply clarify 
-  apply (subgoal_tac "M(<xb,ya>)", blast) 
+  apply (subgoal_tac "M(\<langle>xb,ya\<rangle>)", blast) 
   apply (blast dest: transM) 
  apply auto 
 done
@@ -220,14 +220,14 @@
 done
 
 lemma (in M_basic) wellfounded_on_asym:
-     "\<lbrakk>wellfounded_on(M,A,r);  <a,x>\<in>r;  a\<in>A; x\<in>A;  M(A)\<rbrakk> \<Longrightarrow> <x,a>\<notin>r"
+     "\<lbrakk>wellfounded_on(M,A,r);  \<langle>a,x\<rangle>\<in>r;  a\<in>A; x\<in>A;  M(A)\<rbrakk> \<Longrightarrow> \<langle>x,a\<rangle>\<notin>r"
 apply (simp add: wellfounded_on_def) 
 apply (drule_tac x="{x,a}" in rspec) 
 apply (blast dest: transM)+
 done
 
 lemma (in M_basic) wellordered_asym:
-     "\<lbrakk>wellordered(M,A,r);  <a,x>\<in>r;  a\<in>A; x\<in>A;  M(A)\<rbrakk> \<Longrightarrow> <x,a>\<notin>r"
+     "\<lbrakk>wellordered(M,A,r);  \<langle>a,x\<rangle>\<in>r;  a\<in>A; x\<in>A;  M(A)\<rbrakk> \<Longrightarrow> \<langle>x,a\<rangle>\<notin>r"
 by (simp add: wellordered_def, blast dest: wellfounded_on_asym)
 
 end
--- a/src/ZF/Epsilon.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Epsilon.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,32 +8,32 @@
 theory Epsilon imports Nat begin
 
 definition
-  eclose    :: "i=>i"  where
-    "eclose(A) \<equiv> \<Union>n\<in>nat. nat_rec(n, A, %m r. \<Union>(r))"
+  eclose    :: "i\<Rightarrow>i"  where
+    "eclose(A) \<equiv> \<Union>n\<in>nat. nat_rec(n, A, \<lambda>m r. \<Union>(r))"
 
 definition
-  transrec  :: "[i, [i,i]=>i] =>i"  where
+  transrec  :: "[i, [i,i]\<Rightarrow>i] \<Rightarrow>i"  where
     "transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)"
 
 definition
-  rank      :: "i=>i"  where
-    "rank(a) \<equiv> transrec(a, %x f. \<Union>y\<in>x. succ(f`y))"
+  rank      :: "i\<Rightarrow>i"  where
+    "rank(a) \<equiv> transrec(a, \<lambda>x f. \<Union>y\<in>x. succ(f`y))"
 
 definition
-  transrec2 :: "[i, i, [i,i]=>i] =>i"  where
+  transrec2 :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i"  where
     "transrec2(k, a, b) \<equiv>
        transrec(k,
-                %i r. if(i=0, a,
+                \<lambda>i r. if(i=0, a,
                         if(\<exists>j. i=succ(j),
                            b(THE j. i=succ(j), r`(THE j. i=succ(j))),
                            \<Union>j<i. r`j)))"
 
 definition
-  recursor  :: "[i, [i,i]=>i, i]=>i"  where
-    "recursor(a,b,k) \<equiv>  transrec(k, %n f. nat_case(a, %m. b(m, f`m), n))"
+  recursor  :: "[i, [i,i]\<Rightarrow>i, i]\<Rightarrow>i"  where
+    "recursor(a,b,k) \<equiv>  transrec(k, \<lambda>n f. nat_case(a, \<lambda>m. b(m, f`m), n))"
 
 definition
-  rec  :: "[i, i, [i,i]=>i]=>i"  where
+  rec  :: "[i, i, [i,i]\<Rightarrow>i]\<Rightarrow>i"  where
     "rec(k,a,b) \<equiv> recursor(a,b,k)"
 
 
@@ -85,7 +85,7 @@
 (** eclose(A) is the least transitive set including A as a subset. **)
 
 lemma eclose_least_lemma:
-    "\<lbrakk>Transset(X);  A<=X;  n \<in> nat\<rbrakk> \<Longrightarrow> nat_rec(n, A, %m r. \<Union>(r)) \<subseteq> X"
+    "\<lbrakk>Transset(X);  A<=X;  n \<in> nat\<rbrakk> \<Longrightarrow> nat_rec(n, A, \<lambda>m r. \<Union>(r)) \<subseteq> X"
 apply (unfold Transset_def)
 apply (erule nat_induct)
 apply (simp add: nat_rec_0)
@@ -290,13 +290,13 @@
 apply (erule eclose_rank_lt [THEN succ_leI])
 done
 
-lemma rank_pair1: "rank(a) < rank(<a,b>)"
+lemma rank_pair1: "rank(a) < rank(\<langle>a,b\<rangle>)"
 apply (unfold Pair_def)
 apply (rule consI1 [THEN rank_lt, THEN lt_trans])
 apply (rule consI1 [THEN consI2, THEN rank_lt])
 done
 
-lemma rank_pair2: "rank(b) < rank(<a,b>)"
+lemma rank_pair2: "rank(b) < rank(\<langle>a,b\<rangle>)"
 apply (unfold Pair_def)
 apply (rule consI1 [THEN consI2, THEN rank_lt, THEN lt_trans])
 apply (rule consI1 [THEN consI2, THEN rank_lt])
@@ -310,7 +310,7 @@
 
 (*The first premise not only fixs i but ensures @{term"f\<noteq>0"}.
   The second premise is now essential.  Consider otherwise the relation
-  r = {<0,0>,<0,1>,<0,2>,...}.  Then f`0 = \<Union>(f``{0}) = \<Union>(nat) = nat,
+  r = {\<langle>0,0\<rangle>,\<langle>0,1\<rangle>,\<langle>0,2\<rangle>,...}.  Then f`0 = \<Union>(f``{0}) = \<Union>(nat) = nat,
   whose rank equals that of r.*)
 lemma rank_apply: "\<lbrakk>i \<in> domain(f); function(f)\<rbrakk> \<Longrightarrow> rank(f`i) < rank(f)"
 apply clarify
--- a/src/ZF/EquivClass.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/EquivClass.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,24 +8,24 @@
 theory EquivClass imports Trancl Perm begin
 
 definition
-  quotient   :: "[i,i]=>i"    (infixl \<open>'/'/\<close> 90)  (*set of equiv classes*)  where
+  quotient   :: "[i,i]\<Rightarrow>i"    (infixl \<open>'/'/\<close> 90)  (*set of equiv classes*)  where
       "A//r \<equiv> {r``{x} . x \<in> A}"
 
 definition
-  congruent  :: "[i,i=>i]=>o"  where
-      "congruent(r,b) \<equiv> \<forall>y z. <y,z>:r \<longrightarrow> b(y)=b(z)"
+  congruent  :: "[i,i\<Rightarrow>i]\<Rightarrow>o"  where
+      "congruent(r,b) \<equiv> \<forall>y z. \<langle>y,z\<rangle>:r \<longrightarrow> b(y)=b(z)"
 
 definition
-  congruent2 :: "[i,i,[i,i]=>i]=>o"  where
+  congruent2 :: "[i,i,[i,i]\<Rightarrow>i]\<Rightarrow>o"  where
       "congruent2(r1,r2,b) \<equiv> \<forall>y1 z1 y2 z2.
-           <y1,z1>:r1 \<longrightarrow> <y2,z2>:r2 \<longrightarrow> b(y1,y2) = b(z1,z2)"
+           \<langle>y1,z1\<rangle>:r1 \<longrightarrow> \<langle>y2,z2\<rangle>:r2 \<longrightarrow> b(y1,y2) = b(z1,z2)"
 
 abbreviation
-  RESPECTS ::"[i=>i, i] => o"  (infixr \<open>respects\<close> 80) where
+  RESPECTS ::"[i\<Rightarrow>i, i] \<Rightarrow> o"  (infixr \<open>respects\<close> 80) where
   "f respects r \<equiv> congruent(r,f)"
 
 abbreviation
-  RESPECTS2 ::"[i=>i=>i, i] => o"  (infixr \<open>respects2 \<close> 80) where
+  RESPECTS2 ::"[i\<Rightarrow>i\<Rightarrow>i, i] \<Rightarrow> o"  (infixr \<open>respects2 \<close> 80) where
   "f respects2 r \<equiv> congruent2(r,r,f)"
     \<comment> \<open>Abbreviation for the common case where the relations are identical\<close>
 
@@ -54,18 +54,18 @@
     "\<lbrakk>converse(r) O r = r;  domain(r) = A\<rbrakk> \<Longrightarrow> equiv(A,r)"
 apply (unfold equiv_def refl_def sym_def trans_def)
 apply (erule equalityE)
-apply (subgoal_tac "\<forall>x y. <x,y> \<in> r \<longrightarrow> <y,x> \<in> r", blast+)
+apply (subgoal_tac "\<forall>x y. \<langle>x,y\<rangle> \<in> r \<longrightarrow> \<langle>y,x\<rangle> \<in> r", blast+)
 done
 
 (** Equivalence classes **)
 
 (*Lemma for the next result*)
 lemma equiv_class_subset:
-    "\<lbrakk>sym(r);  trans(r);  <a,b>: r\<rbrakk> \<Longrightarrow> r``{a} \<subseteq> r``{b}"
+    "\<lbrakk>sym(r);  trans(r);  \<langle>a,b\<rangle>: r\<rbrakk> \<Longrightarrow> r``{a} \<subseteq> r``{b}"
 by (unfold trans_def sym_def, blast)
 
 lemma equiv_class_eq:
-    "\<lbrakk>equiv(A,r);  <a,b>: r\<rbrakk> \<Longrightarrow> r``{a} = r``{b}"
+    "\<lbrakk>equiv(A,r);  \<langle>a,b\<rangle>: r\<rbrakk> \<Longrightarrow> r``{a} = r``{b}"
 apply (unfold equiv_def)
 apply (safe del: subsetI intro!: equalityI equiv_class_subset)
 apply (unfold sym_def, blast)
@@ -77,26 +77,26 @@
 
 (*Lemma for the next result*)
 lemma subset_equiv_class:
-    "\<lbrakk>equiv(A,r);  r``{b} \<subseteq> r``{a};  b \<in> A\<rbrakk> \<Longrightarrow> <a,b>: r"
+    "\<lbrakk>equiv(A,r);  r``{b} \<subseteq> r``{a};  b \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle>: r"
 by (unfold equiv_def refl_def, blast)
 
-lemma eq_equiv_class: "\<lbrakk>r``{a} = r``{b};  equiv(A,r);  b \<in> A\<rbrakk> \<Longrightarrow> <a,b>: r"
+lemma eq_equiv_class: "\<lbrakk>r``{a} = r``{b};  equiv(A,r);  b \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle>: r"
 by (assumption | rule equalityD2 subset_equiv_class)+
 
 (*thus r``{a} = r``{b} as well*)
 lemma equiv_class_nondisjoint:
-    "\<lbrakk>equiv(A,r);  x: (r``{a} \<inter> r``{b})\<rbrakk> \<Longrightarrow> <a,b>: r"
+    "\<lbrakk>equiv(A,r);  x: (r``{a} \<inter> r``{b})\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle>: r"
 by (unfold equiv_def trans_def sym_def, blast)
 
 lemma equiv_type: "equiv(A,r) \<Longrightarrow> r \<subseteq> A*A"
 by (unfold equiv_def, blast)
 
 lemma equiv_class_eq_iff:
-     "equiv(A,r) \<Longrightarrow> <x,y>: r \<longleftrightarrow> r``{x} = r``{y} \<and> x \<in> A \<and> y \<in> A"
+     "equiv(A,r) \<Longrightarrow> \<langle>x,y\<rangle>: r \<longleftrightarrow> r``{x} = r``{y} \<and> x \<in> A \<and> y \<in> A"
 by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)
 
 lemma eq_equiv_class_iff:
-     "\<lbrakk>equiv(A,r);  x \<in> A;  y \<in> A\<rbrakk> \<Longrightarrow> r``{x} = r``{y} \<longleftrightarrow> <x,y>: r"
+     "\<lbrakk>equiv(A,r);  x \<in> A;  y \<in> A\<rbrakk> \<Longrightarrow> r``{x} = r``{y} \<longleftrightarrow> \<langle>x,y\<rangle>: r"
 by (blast intro: eq_equiv_class equiv_class_eq dest: equiv_type)
 
 (*** Quotients ***)
@@ -151,7 +151,7 @@
 lemma UN_equiv_class_inject:
     "\<lbrakk>equiv(A,r);   b respects r;
         (\<Union>x\<in>X. b(x))=(\<Union>y\<in>Y. b(y));  X \<in> A//r;  Y \<in> A//r;
-        \<And>x y. \<lbrakk>x \<in> A; y \<in> A; b(x)=b(y)\<rbrakk> \<Longrightarrow> <x,y>:r\<rbrakk>
+        \<And>x y. \<lbrakk>x \<in> A; y \<in> A; b(x)=b(y)\<rbrakk> \<Longrightarrow> \<langle>x,y\<rangle>:r\<rbrakk>
      \<Longrightarrow> X=Y"
 apply (unfold quotient_def, safe)
 apply (rule equiv_class_eq, assumption)
@@ -167,7 +167,7 @@
 
 lemma congruent2_implies_congruent_UN:
     "\<lbrakk>equiv(A1,r1);  equiv(A2,r2);  congruent2(r1,r2,b);  a \<in> A2\<rbrakk> \<Longrightarrow>
-     congruent(r1, %x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))"
+     congruent(r1, \<lambda>x1. \<Union>x2 \<in> r2``{a}. b(x1,x2))"
 apply (unfold congruent_def, safe)
 apply (frule equiv_type [THEN subsetD], assumption)
 apply clarify
@@ -197,8 +197,8 @@
   than the direct proof*)
 lemma congruent2I:
     "\<lbrakk>equiv(A1,r1);  equiv(A2,r2);
-        \<And>y z w. \<lbrakk>w \<in> A2;  <y,z> \<in> r1\<rbrakk> \<Longrightarrow> b(y,w) = b(z,w);
-        \<And>y z w. \<lbrakk>w \<in> A1;  <y,z> \<in> r2\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)
+        \<And>y z w. \<lbrakk>w \<in> A2;  \<langle>y,z\<rangle> \<in> r1\<rbrakk> \<Longrightarrow> b(y,w) = b(z,w);
+        \<And>y z w. \<lbrakk>w \<in> A1;  \<langle>y,z\<rangle> \<in> r2\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)
 \<rbrakk> \<Longrightarrow> congruent2(r1,r2,b)"
 apply (unfold congruent2_def equiv_def refl_def, safe)
 apply (blast intro: trans)
@@ -207,7 +207,7 @@
 lemma congruent2_commuteI:
  assumes equivA: "equiv(A,r)"
      and commute: "\<And>y z. \<lbrakk>y \<in> A;  z \<in> A\<rbrakk> \<Longrightarrow> b(y,z) = b(z,y)"
-     and congt:   "\<And>y z w. \<lbrakk>w \<in> A;  <y,z>: r\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)"
+     and congt:   "\<And>y z w. \<lbrakk>w \<in> A;  \<langle>y,z\<rangle>: r\<rbrakk> \<Longrightarrow> b(w,y) = b(w,z)"
  shows "b respects2 r"
 apply (insert equivA [THEN equiv_type, THEN subsetD])
 apply (rule congruent2I [OF equivA equivA])
@@ -220,9 +220,9 @@
 (*Obsolete?*)
 lemma congruent_commuteI:
     "\<lbrakk>equiv(A,r);  Z \<in> A//r;
-        \<And>w. \<lbrakk>w \<in> A\<rbrakk> \<Longrightarrow> congruent(r, %z. b(w,z));
+        \<And>w. \<lbrakk>w \<in> A\<rbrakk> \<Longrightarrow> congruent(r, \<lambda>z. b(w,z));
         \<And>x y. \<lbrakk>x \<in> A;  y \<in> A\<rbrakk> \<Longrightarrow> b(y,x) = b(x,y)
-\<rbrakk> \<Longrightarrow> congruent(r, %w. \<Union>z\<in>Z. b(w,z))"
+\<rbrakk> \<Longrightarrow> congruent(r, \<lambda>w. \<Union>z\<in>Z. b(w,z))"
 apply (simp (no_asm) add: congruent_def)
 apply (safe elim!: quotientE)
 apply (frule equiv_type [THEN subsetD], assumption)
--- a/src/ZF/Finite.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Finite.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -18,8 +18,8 @@
 
 
 consts
-  Fin       :: "i=>i"
-  FiniteFun :: "[i,i]=>i"         (\<open>(_ -||>/ _)\<close> [61, 60] 60)
+  Fin       :: "i\<Rightarrow>i"
+  FiniteFun :: "[i,i]\<Rightarrow>i"         (\<open>(_ -||>/ _)\<close> [61, 60] 60)
 
 inductive
   domains   "Fin(A)" \<subseteq> "Pow(A)"
@@ -34,7 +34,7 @@
   intros
     emptyI:  "0 \<in> A -||> B"
     consI:   "\<lbrakk>a \<in> A;  b \<in> B;  h \<in> A -||> B;  a \<notin> domain(h)\<rbrakk>
-              \<Longrightarrow> cons(<a,b>,h) \<in> A -||> B"
+              \<Longrightarrow> cons(\<langle>a,b\<rangle>,h) \<in> A -||> B"
   type_intros Fin.intros
 
 
@@ -200,7 +200,7 @@
 subsection\<open>The Contents of a Singleton Set\<close>
 
 definition
-  contents :: "i=>i"  where
+  contents :: "i\<Rightarrow>i"  where
    "contents(X) \<equiv> THE x. X = {x}"
 
 lemma contents_eq [simp]: "contents ({x}) = x"
--- a/src/ZF/Fixedpt.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Fixedpt.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,15 +9,15 @@
 
 definition 
   (*monotone operator from Pow(D) to itself*)
-  bnd_mono :: "[i,i=>i]=>o"  where
+  bnd_mono :: "[i,i\<Rightarrow>i]\<Rightarrow>o"  where
      "bnd_mono(D,h) \<equiv> h(D)<=D \<and> (\<forall>W X. W<=X \<longrightarrow> X<=D \<longrightarrow> h(W) \<subseteq> h(X))"
 
 definition 
-  lfp      :: "[i,i=>i]=>i"  where
+  lfp      :: "[i,i\<Rightarrow>i]\<Rightarrow>i"  where
      "lfp(D,h) \<equiv> \<Inter>({X: Pow(D). h(X) \<subseteq> X})"
 
 definition 
-  gfp      :: "[i,i=>i]=>i"  where
+  gfp      :: "[i,i\<Rightarrow>i]\<Rightarrow>i"  where
      "gfp(D,h) \<equiv> \<Union>({X: Pow(D). X \<subseteq> h(X)})"
 
 text\<open>The theorem is proved in the lattice of subsets of \<^term>\<open>D\<close>, 
@@ -285,7 +285,7 @@
 
 (*The version used in the induction/coinduction package*)
 lemma def_Collect_coinduct:
-    "\<lbrakk>A \<equiv> gfp(D, %w. Collect(D,P(w)));  bnd_mono(D, %w. Collect(D,P(w)));   
+    "\<lbrakk>A \<equiv> gfp(D, \<lambda>w. Collect(D,P(w)));  bnd_mono(D, \<lambda>w. Collect(D,P(w)));   
         a: X;  X \<subseteq> D;  \<And>z. z: X \<Longrightarrow> P(X \<union> A, z)\<rbrakk> \<Longrightarrow>  
      a \<in> A"
 apply (rule def_coinduct, assumption+, blast+)
--- a/src/ZF/IMP/Com.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/IMP/Com.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -23,17 +23,17 @@
 consts evala :: i
 
 abbreviation
-  evala_syntax :: "[i, i] => o"    (infixl \<open>-a->\<close> 50)
-  where "p -a-> n \<equiv> <p,n> \<in> evala"
+  evala_syntax :: "[i, i] \<Rightarrow> o"    (infixl \<open>-a->\<close> 50)
+  where "p -a-> n \<equiv> \<langle>p,n\<rangle> \<in> evala"
 
 inductive
   domains "evala" \<subseteq> "(aexp \<times> (loc -> nat)) \<times> nat"
   intros
     N:   "\<lbrakk>n \<in> nat;  sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> <N(n),sigma> -a-> n"
     X:   "\<lbrakk>x \<in> loc;  sigma \<in> loc->nat\<rbrakk> \<Longrightarrow> <X(x),sigma> -a-> sigma`x"
-    Op1: "\<lbrakk><e,sigma> -a-> n; f \<in> nat -> nat\<rbrakk> \<Longrightarrow> <Op1(f,e),sigma> -a-> f`n"
-    Op2: "\<lbrakk><e0,sigma> -a-> n0;  <e1,sigma>  -a-> n1; f \<in> (nat\<times>nat) -> nat\<rbrakk>
-          \<Longrightarrow> <Op2(f,e0,e1),sigma> -a-> f`<n0,n1>"
+    Op1: "\<lbrakk>\<langle>e,sigma\<rangle> -a-> n; f \<in> nat -> nat\<rbrakk> \<Longrightarrow> <Op1(f,e),sigma> -a-> f`n"
+    Op2: "\<lbrakk>\<langle>e0,sigma\<rangle> -a-> n0;  \<langle>e1,sigma\<rangle>  -a-> n1; f \<in> (nat\<times>nat) -> nat\<rbrakk>
+          \<Longrightarrow> <Op2(f,e0,e1),sigma> -a-> f`\<langle>n0,n1\<rangle>"
   type_intros aexp.intros apply_funtype
 
 
@@ -53,20 +53,20 @@
 consts evalb :: i
 
 abbreviation
-  evalb_syntax :: "[i,i] => o"    (infixl \<open>-b->\<close> 50)
-  where "p -b-> b \<equiv> <p,b> \<in> evalb"
+  evalb_syntax :: "[i,i] \<Rightarrow> o"    (infixl \<open>-b->\<close> 50)
+  where "p -b-> b \<equiv> \<langle>p,b\<rangle> \<in> evalb"
 
 inductive
   domains "evalb" \<subseteq> "(bexp \<times> (loc -> nat)) \<times> bool"
   intros
-    true:  "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> <true,sigma> -b-> 1"
-    false: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> <false,sigma> -b-> 0"
-    ROp:   "\<lbrakk><a0,sigma> -a-> n0; <a1,sigma> -a-> n1; f \<in> (nat*nat)->bool\<rbrakk>
-           \<Longrightarrow> <ROp(f,a0,a1),sigma> -b-> f`<n0,n1> "
-    noti:  "\<lbrakk><b,sigma> -b-> w\<rbrakk> \<Longrightarrow> <noti(b),sigma> -b-> not(w)"
-    andi:  "\<lbrakk><b0,sigma> -b-> w0; <b1,sigma> -b-> w1\<rbrakk>
+    true:  "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> \<langle>true,sigma\<rangle> -b-> 1"
+    false: "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> \<langle>false,sigma\<rangle> -b-> 0"
+    ROp:   "\<lbrakk>\<langle>a0,sigma\<rangle> -a-> n0; \<langle>a1,sigma\<rangle> -a-> n1; f \<in> (nat*nat)->bool\<rbrakk>
+           \<Longrightarrow> <ROp(f,a0,a1),sigma> -b-> f`\<langle>n0,n1\<rangle> "
+    noti:  "\<lbrakk>\<langle>b,sigma\<rangle> -b-> w\<rbrakk> \<Longrightarrow> <noti(b),sigma> -b-> not(w)"
+    andi:  "\<lbrakk>\<langle>b0,sigma\<rangle> -b-> w0; \<langle>b1,sigma\<rangle> -b-> w1\<rbrakk>
           \<Longrightarrow> <b0 andi b1,sigma> -b-> (w0 and w1)"
-    ori:   "\<lbrakk><b0,sigma> -b-> w0; <b1,sigma> -b-> w1\<rbrakk>
+    ori:   "\<lbrakk>\<langle>b0,sigma\<rangle> -b-> w0; \<langle>b1,sigma\<rangle> -b-> w1\<rbrakk>
             \<Longrightarrow> <b0 ori b1,sigma> -b-> (w0 or w1)"
   type_intros  bexp.intros
                apply_funtype and_type or_type bool_1I bool_0I not_type
@@ -87,32 +87,32 @@
 consts evalc :: i
 
 abbreviation
-  evalc_syntax :: "[i, i] => o"    (infixl \<open>-c->\<close> 50)
-  where "p -c-> s \<equiv> <p,s> \<in> evalc"
+  evalc_syntax :: "[i, i] \<Rightarrow> o"    (infixl \<open>-c->\<close> 50)
+  where "p -c-> s \<equiv> \<langle>p,s\<rangle> \<in> evalc"
 
 inductive
   domains "evalc" \<subseteq> "(com \<times> (loc -> nat)) \<times> (loc -> nat)"
   intros
     skip:    "\<lbrakk>sigma \<in> loc -> nat\<rbrakk> \<Longrightarrow> <\<SKIP>,sigma> -c-> sigma"
 
-    assign:  "\<lbrakk>m \<in> nat; x \<in> loc; <a,sigma> -a-> m\<rbrakk>
+    assign:  "\<lbrakk>m \<in> nat; x \<in> loc; \<langle>a,sigma\<rangle> -a-> m\<rbrakk>
               \<Longrightarrow> <x \<ASSN> a,sigma> -c-> sigma(x:=m)"
 
-    semi:    "\<lbrakk><c0,sigma> -c-> sigma2; <c1,sigma2> -c-> sigma1\<rbrakk>
+    semi:    "\<lbrakk>\<langle>c0,sigma\<rangle> -c-> sigma2; \<langle>c1,sigma2\<rangle> -c-> sigma1\<rbrakk>
               \<Longrightarrow> <c0\<SEQ> c1, sigma> -c-> sigma1"
 
     if1:     "\<lbrakk>b \<in> bexp; c1 \<in> com; sigma \<in> loc->nat;
-                 <b,sigma> -b-> 1; <c0,sigma> -c-> sigma1\<rbrakk>
+                 \<langle>b,sigma\<rangle> -b-> 1; \<langle>c0,sigma\<rangle> -c-> sigma1\<rbrakk>
               \<Longrightarrow> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1"
 
     if0:     "\<lbrakk>b \<in> bexp; c0 \<in> com; sigma \<in> loc->nat;
-                 <b,sigma> -b-> 0; <c1,sigma> -c-> sigma1\<rbrakk>
+                 \<langle>b,sigma\<rangle> -b-> 0; \<langle>c1,sigma\<rangle> -c-> sigma1\<rbrakk>
                \<Longrightarrow> <\<IF> b \<THEN> c0 \<ELSE> c1, sigma> -c-> sigma1"
 
-    while0:   "\<lbrakk>c \<in> com; <b, sigma> -b-> 0\<rbrakk>
+    while0:   "\<lbrakk>c \<in> com; \<langle>b, sigma\<rangle> -b-> 0\<rbrakk>
                \<Longrightarrow> <\<WHILE> b \<DO> c,sigma> -c-> sigma"
 
-    while1:   "\<lbrakk>c \<in> com; <b,sigma> -b-> 1; <c,sigma> -c-> sigma2;
+    while1:   "\<lbrakk>c \<in> com; \<langle>b,sigma\<rangle> -b-> 1; \<langle>c,sigma\<rangle> -c-> sigma2;
                   <\<WHILE> b \<DO> c, sigma2> -c-> sigma1\<rbrakk>
                \<Longrightarrow> <\<WHILE> b \<DO> c, sigma> -c-> sigma1"
 
--- a/src/ZF/IMP/Denotation.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/IMP/Denotation.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,12 +9,12 @@
 subsection \<open>Definitions\<close>
 
 consts
-  A     :: "i => i => i"
-  B     :: "i => i => i"
-  C     :: "i => i"
+  A     :: "i \<Rightarrow> i \<Rightarrow> i"
+  B     :: "i \<Rightarrow> i \<Rightarrow> i"
+  C     :: "i \<Rightarrow> i"
 
 definition
-  Gamma :: "[i,i,i] => i"  (\<open>\<Gamma>\<close>) where
+  Gamma :: "[i,i,i] \<Rightarrow> i"  (\<open>\<Gamma>\<close>) where
   "\<Gamma>(b,cden) \<equiv>
     (\<lambda>phi. {io \<in> (phi O cden). B(b,fst(io))=1} \<union>
            {io \<in> id(loc->nat). B(b,fst(io))=0})"
@@ -58,7 +58,7 @@
   done
 
 lemma C_type_D [dest]:
-    "\<lbrakk><x,y> \<in> C(c); c \<in> com\<rbrakk> \<Longrightarrow> x \<in> loc->nat \<and> y \<in> loc->nat"
+    "\<lbrakk>\<langle>x,y\<rangle> \<in> C(c); c \<in> com\<rbrakk> \<Longrightarrow> x \<in> loc->nat \<and> y \<in> loc->nat"
   by (blast dest: C_subset [THEN subsetD])
 
 lemma C_type_fst [dest]: "\<lbrakk>x \<in> C(c); c \<in> com\<rbrakk> \<Longrightarrow> fst(x) \<in> loc->nat"
--- a/src/ZF/IMP/Equiv.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/IMP/Equiv.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,7 +8,7 @@
 
 lemma aexp_iff [rule_format]:
   "\<lbrakk>a \<in> aexp; sigma: loc -> nat\<rbrakk> 
-    \<Longrightarrow> \<forall>n. <a,sigma> -a-> n \<longleftrightarrow> A(a,sigma) = n"
+    \<Longrightarrow> \<forall>n. \<langle>a,sigma\<rangle> -a-> n \<longleftrightarrow> A(a,sigma) = n"
   apply (erule aexp.induct)
      apply (force intro!: evala.intros)+
   done
@@ -17,8 +17,8 @@
         aexp_iff [THEN iffD2, intro!]
 
 inductive_cases [elim!]:
-  "<true,sigma> -b-> x"
-  "<false,sigma> -b-> x"
+  "\<langle>true,sigma\<rangle> -b-> x"
+  "\<langle>false,sigma\<rangle> -b-> x"
   "<ROp(f,a0,a1),sigma> -b-> x"
   "<noti(b),sigma> -b-> x"
   "<b0 andi b1,sigma> -b-> x"
@@ -27,7 +27,7 @@
 
 lemma bexp_iff [rule_format]:
   "\<lbrakk>b \<in> bexp; sigma: loc -> nat\<rbrakk> 
-    \<Longrightarrow> \<forall>w. <b,sigma> -b-> w \<longleftrightarrow> B(b,sigma) = w"
+    \<Longrightarrow> \<forall>w. \<langle>b,sigma\<rangle> -b-> w \<longleftrightarrow> B(b,sigma) = w"
   apply (erule bexp.induct) 
   apply (auto intro!: evalb.intros)
   done
@@ -35,7 +35,7 @@
 declare bexp_iff [THEN iffD1, simp]
         bexp_iff [THEN iffD2, intro!]
 
-lemma com1: "<c,sigma> -c-> sigma' \<Longrightarrow> <sigma,sigma'> \<in> C(c)"
+lemma com1: "\<langle>c,sigma\<rangle> -c-> sigma' \<Longrightarrow> <sigma,sigma'> \<in> C(c)"
   apply (erule evalc.induct)
         apply (simp_all (no_asm_simp))
      txt \<open>\<open>assign\<close>\<close>
--- a/src/ZF/Induct/Acc.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Acc.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,7 +12,7 @@
 \<close>
 
 consts
-  acc :: "i => i"
+  acc :: "i \<Rightarrow> i"
 
 inductive
   domains "acc(r)" \<subseteq> "field(r)"
@@ -28,15 +28,15 @@
   The intended introduction rule:
 \<close>
 
-lemma accI: "\<lbrakk>\<And>b. <b,a>:r \<Longrightarrow> b \<in> acc(r);  a \<in> field(r)\<rbrakk> \<Longrightarrow> a \<in> acc(r)"
+lemma accI: "\<lbrakk>\<And>b. \<langle>b,a\<rangle>:r \<Longrightarrow> b \<in> acc(r);  a \<in> field(r)\<rbrakk> \<Longrightarrow> a \<in> acc(r)"
   by (blast intro: acc.intros)
 
-lemma acc_downward: "\<lbrakk>b \<in> acc(r);  <a,b>: r\<rbrakk> \<Longrightarrow> a \<in> acc(r)"
+lemma acc_downward: "\<lbrakk>b \<in> acc(r);  \<langle>a,b\<rangle>: r\<rbrakk> \<Longrightarrow> a \<in> acc(r)"
   by (erule acc.cases) blast
 
 lemma acc_induct [consumes 1, case_names vimage, induct set: acc]:
     "\<lbrakk>a \<in> acc(r);
-        \<And>x. \<lbrakk>x \<in> acc(r);  \<forall>y. <y,x>:r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)
+        \<And>x. \<lbrakk>x \<in> acc(r);  \<forall>y. \<langle>y,x\<rangle>:r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)
 \<rbrakk> \<Longrightarrow> P(a)"
   by (erule acc.induct) (blast intro: acc.intros)
 
--- a/src/ZF/Induct/Binary_Trees.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Binary_Trees.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,7 +10,7 @@
 subsection \<open>Datatype definition\<close>
 
 consts
-  bt :: "i => i"
+  bt :: "i \<Rightarrow> i"
 
 datatype "bt(A)" =
   Lf | Br ("a \<in> A", "t1 \<in> bt(A)", "t2 \<in> bt(A)")
@@ -66,7 +66,7 @@
 
 subsection \<open>Number of nodes, with an example of tail-recursion\<close>
 
-consts  n_nodes :: "i => i"
+consts  n_nodes :: "i \<Rightarrow> i"
 primrec
   "n_nodes(Lf) = 0"
   "n_nodes(Br(a, l, r)) = succ(n_nodes(l) #+ n_nodes(r))"
@@ -74,7 +74,7 @@
 lemma n_nodes_type [simp]: "t \<in> bt(A) \<Longrightarrow> n_nodes(t) \<in> nat"
   by (induct set: bt) auto
 
-consts  n_nodes_aux :: "i => i"
+consts  n_nodes_aux :: "i \<Rightarrow> i"
 primrec
   "n_nodes_aux(Lf) = (\<lambda>k \<in> nat. k)"
   "n_nodes_aux(Br(a, l, r)) =
@@ -88,7 +88,7 @@
   done
 
 definition
-  n_nodes_tail :: "i => i"  where
+  n_nodes_tail :: "i \<Rightarrow> i"  where
   "n_nodes_tail(t) \<equiv> n_nodes_aux(t) ` 0"
 
 lemma "t \<in> bt(A) \<Longrightarrow> n_nodes_tail(t) = n_nodes(t)"
@@ -98,7 +98,7 @@
 subsection \<open>Number of leaves\<close>
 
 consts
-  n_leaves :: "i => i"
+  n_leaves :: "i \<Rightarrow> i"
 primrec
   "n_leaves(Lf) = 1"
   "n_leaves(Br(a, l, r)) = n_leaves(l) #+ n_leaves(r)"
@@ -110,7 +110,7 @@
 subsection \<open>Reflecting trees\<close>
 
 consts
-  bt_reflect :: "i => i"
+  bt_reflect :: "i \<Rightarrow> i"
 primrec
   "bt_reflect(Lf) = Lf"
   "bt_reflect(Br(a, l, r)) = Br(a, bt_reflect(r), bt_reflect(l))"
--- a/src/ZF/Induct/Brouwer.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Brouwer.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -42,7 +42,7 @@
 subsection \<open>The Martin-Löf wellordering type\<close>
 
 consts
-  Well :: "[i, i => i] => i"
+  Well :: "[i, i \<Rightarrow> i] \<Rightarrow> i"
 
 datatype \<subseteq> "Vfrom(A \<union> (\<Union>x \<in> A. B(x)), csucc(nat \<union> |\<Union>x \<in> A. B(x)|))"
     \<comment> \<open>The union with \<open>nat\<close> ensures that the cardinal is infinite.\<close>
--- a/src/ZF/Induct/Comb.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Comb.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -33,10 +33,10 @@
 
 consts contract  :: i
 abbreviation contract_syntax :: "[i,i] \<Rightarrow> o"  (infixl \<open>\<rightarrow>\<^sup>1\<close> 50)
-  where "p \<rightarrow>\<^sup>1 q \<equiv> <p,q> \<in> contract"
+  where "p \<rightarrow>\<^sup>1 q \<equiv> \<langle>p,q\<rangle> \<in> contract"
 
 abbreviation contract_multi :: "[i,i] \<Rightarrow> o"  (infixl \<open>\<rightarrow>\<close> 50)
-  where "p \<rightarrow> q \<equiv> <p,q> \<in> contract^*"
+  where "p \<rightarrow> q \<equiv> \<langle>p,q\<rangle> \<in> contract^*"
 
 inductive
   domains "contract" \<subseteq> "comb \<times> comb"
@@ -54,11 +54,11 @@
 
 consts parcontract :: i
 
-abbreviation parcontract_syntax :: "[i,i] => o"  (infixl \<open>\<Rrightarrow>\<^sup>1\<close> 50)
-  where "p \<Rrightarrow>\<^sup>1 q \<equiv> <p,q> \<in> parcontract"
+abbreviation parcontract_syntax :: "[i,i] \<Rightarrow> o"  (infixl \<open>\<Rrightarrow>\<^sup>1\<close> 50)
+  where "p \<Rrightarrow>\<^sup>1 q \<equiv> \<langle>p,q\<rangle> \<in> parcontract"
 
-abbreviation parcontract_multi :: "[i,i] => o"  (infixl \<open>\<Rrightarrow>\<close> 50)
-  where "p \<Rrightarrow> q \<equiv> <p,q> \<in> parcontract^+"
+abbreviation parcontract_multi :: "[i,i] \<Rightarrow> o"  (infixl \<open>\<Rrightarrow>\<close> 50)
+  where "p \<Rrightarrow> q \<equiv> \<langle>p,q\<rangle> \<in> parcontract^+"
 
 inductive
   domains "parcontract" \<subseteq> "comb \<times> comb"
@@ -78,14 +78,14 @@
 
 definition diamond :: "i \<Rightarrow> o"
   where "diamond(r) \<equiv>
-    \<forall>x y. <x,y>\<in>r \<longrightarrow> (\<forall>y'. <x,y'>\<in>r \<longrightarrow> (\<exists>z. <y,z>\<in>r \<and> <y',z> \<in> r))"
+    \<forall>x y. \<langle>x,y\<rangle>\<in>r \<longrightarrow> (\<forall>y'. <x,y'>\<in>r \<longrightarrow> (\<exists>z. \<langle>y,z\<rangle>\<in>r \<and> <y',z> \<in> r))"
 
 
 subsection \<open>Transitive closure preserves the Church-Rosser property\<close>
 
 lemma diamond_strip_lemmaD [rule_format]:
-  "\<lbrakk>diamond(r);  <x,y>:r^+\<rbrakk> \<Longrightarrow>
-    \<forall>y'. <x,y'>:r \<longrightarrow> (\<exists>z. <y',z>: r^+ \<and> <y,z>: r)"
+  "\<lbrakk>diamond(r);  \<langle>x,y\<rangle>:r^+\<rbrakk> \<Longrightarrow>
+    \<forall>y'. <x,y'>:r \<longrightarrow> (\<exists>z. <y',z>: r^+ \<and> \<langle>y,z\<rangle>: r)"
   apply (unfold diamond_def)
   apply (erule trancl_induct)
    apply (blast intro: r_into_trancl)
--- a/src/ZF/Induct/Datatypes.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Datatypes.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -14,7 +14,7 @@
 \<close>
 
 consts
-  data :: "[i, i] => i"
+  data :: "[i, i] \<Rightarrow> i"
 
 datatype "data(A, B)" =
     Con0
--- a/src/ZF/Induct/FoldSet.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/FoldSet.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,28 +8,28 @@
 
 theory FoldSet imports ZF begin
 
-consts fold_set :: "[i, i, [i,i]=>i, i] => i"
+consts fold_set :: "[i, i, [i,i]\<Rightarrow>i, i] \<Rightarrow> i"
 
 inductive
   domains "fold_set(A, B, f,e)" \<subseteq> "Fin(A)*B"
   intros
-    emptyI: "e\<in>B \<Longrightarrow> <0, e>\<in>fold_set(A, B, f,e)"
-    consI:  "\<lbrakk>x\<in>A; x \<notin>C;  <C,y> \<in> fold_set(A, B,f,e); f(x,y):B\<rbrakk>
+    emptyI: "e\<in>B \<Longrightarrow> \<langle>0, e\<rangle>\<in>fold_set(A, B, f,e)"
+    consI:  "\<lbrakk>x\<in>A; x \<notin>C;  \<langle>C,y\<rangle> \<in> fold_set(A, B,f,e); f(x,y):B\<rbrakk>
                 \<Longrightarrow>  <cons(x,C), f(x,y)>\<in>fold_set(A, B, f, e)"
   type_intros Fin.intros
   
 definition
-  fold :: "[i, [i,i]=>i, i, i] => i"  (\<open>fold[_]'(_,_,_')\<close>)  where
-   "fold[B](f,e, A) \<equiv> THE x. <A, x>\<in>fold_set(A, B, f,e)"
+  fold :: "[i, [i,i]\<Rightarrow>i, i, i] \<Rightarrow> i"  (\<open>fold[_]'(_,_,_')\<close>)  where
+   "fold[B](f,e, A) \<equiv> THE x. \<langle>A, x\<rangle>\<in>fold_set(A, B, f,e)"
 
 definition
-   setsum :: "[i=>i, i] => i"  where
+   setsum :: "[i\<Rightarrow>i, i] \<Rightarrow> i"  where
    "setsum(g, C) \<equiv> if Finite(C) then
-                     fold[int](%x y. g(x) $+ y, #0, C) else #0"
+                     fold[int](\<lambda>x y. g(x) $+ y, #0, C) else #0"
 
 (** foldSet **)
 
-inductive_cases empty_fold_setE: "<0, x> \<in> fold_set(A, B, f,e)"
+inductive_cases empty_fold_setE: "\<langle>0, x\<rangle> \<in> fold_set(A, B, f,e)"
 inductive_cases cons_fold_setE: "<cons(x,C), y> \<in> fold_set(A, B, f,e)"
 
 (* add-hoc lemmas *)                                                
@@ -44,8 +44,8 @@
 
 (* fold_set monotonicity *)
 lemma fold_set_mono_lemma:
-     "<C, x> \<in> fold_set(A, B, f, e)  
-      \<Longrightarrow> \<forall>D. A<=D \<longrightarrow> <C, x> \<in> fold_set(D, B, f, e)"
+     "\<langle>C, x\<rangle> \<in> fold_set(A, B, f, e)  
+      \<Longrightarrow> \<forall>D. A<=D \<longrightarrow> \<langle>C, x\<rangle> \<in> fold_set(D, B, f, e)"
 apply (erule fold_set.induct)
 apply (auto intro: fold_set.intros)
 done
@@ -57,7 +57,7 @@
 done
 
 lemma fold_set_lemma:
-     "<C, x>\<in>fold_set(A, B, f, e) \<Longrightarrow> <C, x>\<in>fold_set(C, B, f, e) \<and> C<=A"
+     "\<langle>C, x\<rangle>\<in>fold_set(A, B, f, e) \<Longrightarrow> \<langle>C, x\<rangle>\<in>fold_set(C, B, f, e) \<and> C<=A"
 apply (erule fold_set.induct)
 apply (auto intro!: fold_set.intros intro: fold_set_mono [THEN subsetD])
 done
@@ -79,7 +79,7 @@
 
 
 lemma (in fold_typing) Fin_imp_fold_set:
-     "C\<in>Fin(A) \<Longrightarrow> (\<exists>x. <C, x> \<in> fold_set(A, B, f,e))"
+     "C\<in>Fin(A) \<Longrightarrow> (\<exists>x. \<langle>C, x\<rangle> \<in> fold_set(A, B, f,e))"
 apply (erule Fin_induct)
 apply (auto dest: fold_set.dom_subset [THEN subsetD] 
             intro: fold_set.intros etype ftype)
@@ -92,8 +92,8 @@
 lemma (in fold_typing) fold_set_determ_lemma [rule_format]: 
 "n\<in>nat
  \<Longrightarrow> \<forall>C. |C|<n \<longrightarrow>  
-   (\<forall>x. <C, x> \<in> fold_set(A, B, f,e)\<longrightarrow> 
-           (\<forall>y. <C, y> \<in> fold_set(A, B, f,e) \<longrightarrow> y=x))"
+   (\<forall>x. \<langle>C, x\<rangle> \<in> fold_set(A, B, f,e)\<longrightarrow> 
+           (\<forall>y. \<langle>C, y\<rangle> \<in> fold_set(A, B, f,e) \<longrightarrow> y=x))"
 apply (erule nat_induct)
  apply (auto simp add: le_iff)
 apply (erule fold_set.cases)
@@ -131,8 +131,8 @@
 done
 
 lemma (in fold_typing) fold_set_determ: 
-     "\<lbrakk><C, x>\<in>fold_set(A, B, f, e);  
-         <C, y>\<in>fold_set(A, B, f, e)\<rbrakk> \<Longrightarrow> y=x"
+     "\<lbrakk>\<langle>C, x\<rangle>\<in>fold_set(A, B, f, e);  
+         \<langle>C, y\<rangle>\<in>fold_set(A, B, f, e)\<rbrakk> \<Longrightarrow> y=x"
 apply (frule fold_set.dom_subset [THEN subsetD], clarify)
 apply (drule Fin_into_Finite)
 apply (unfold Finite_def, clarify)
@@ -143,7 +143,7 @@
 (** The fold function **)
 
 lemma (in fold_typing) fold_equality: 
-     "<C,y> \<in> fold_set(A,B,f,e) \<Longrightarrow> fold[B](f,e,C) = y"
+     "\<langle>C,y\<rangle> \<in> fold_set(A,B,f,e) \<Longrightarrow> fold[B](f,e,C) = y"
 apply (unfold fold_def)
 apply (frule fold_set.dom_subset [THEN subsetD], clarify)
 apply (rule the_equality)
@@ -161,7 +161,7 @@
 
 text\<open>This result is the right-to-left direction of the subsequent result\<close>
 lemma (in fold_typing) fold_set_imp_cons: 
-     "\<lbrakk><C, y> \<in> fold_set(C, B, f, e); C \<in> Fin(A); c \<in> A; c\<notin>C\<rbrakk>
+     "\<lbrakk>\<langle>C, y\<rangle> \<in> fold_set(C, B, f, e); C \<in> Fin(A); c \<in> A; c\<notin>C\<rbrakk>
       \<Longrightarrow> <cons(c, C), f(c,y)> \<in> fold_set(cons(c, C), B, f, e)"
 apply (frule FinD [THEN fold_set_mono, THEN subsetD])
  apply assumption
@@ -172,7 +172,7 @@
 lemma (in fold_typing) fold_cons_lemma [rule_format]: 
 "\<lbrakk>C \<in> Fin(A); c \<in> A; c\<notin>C\<rbrakk>   
      \<Longrightarrow> <cons(c, C), v> \<in> fold_set(cons(c, C), B, f, e) \<longleftrightarrow>   
-         (\<exists>y. <C, y> \<in> fold_set(C, B, f, e) \<and> v = f(c, y))"
+         (\<exists>y. \<langle>C, y\<rangle> \<in> fold_set(C, B, f, e) \<and> v = f(c, y))"
 apply auto
  prefer 2 apply (blast intro: fold_set_imp_cons) 
  apply (frule_tac Fin.consI [of c, THEN FinD, THEN fold_set_mono, THEN subsetD], assumption+)
@@ -251,7 +251,7 @@
 apply (auto intro: fold_typing.intro Finite_cons_lemma)
 done
 
-lemma setsum_K0: "setsum((%i. #0), C) = #0"
+lemma setsum_K0: "setsum((\<lambda>i. #0), C) = #0"
 apply (case_tac "Finite (C) ")
  prefer 2 apply (simp add: setsum_def)
 apply (erule Finite_induct, auto)
@@ -290,7 +290,7 @@
      "Finite(I)  
       \<Longrightarrow> (\<forall>i\<in>I. Finite(C(i))) \<longrightarrow>  
           (\<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j \<longrightarrow> C(i) \<inter> C(j) = 0) \<longrightarrow>  
-          setsum(f, \<Union>i\<in>I. C(i)) = setsum (%i. setsum(f, C(i)), I)"
+          setsum(f, \<Union>i\<in>I. C(i)) = setsum (\<lambda>i. setsum(f, C(i)), I)"
 apply (erule Finite_induct, auto)
 apply (subgoal_tac "\<forall>i\<in>B. x \<noteq> i")
  prefer 2 apply blast
@@ -302,7 +302,7 @@
 done
 
 
-lemma setsum_addf: "setsum(%x. f(x) $+ g(x),C) = setsum(f, C) $+ setsum(g, C)"
+lemma setsum_addf: "setsum(\<lambda>x. f(x) $+ g(x),C) = setsum(f, C) $+ setsum(g, C)"
 apply (case_tac "Finite (C) ")
  prefer 2 apply (simp add: setsum_def)
 apply (erule Finite_induct, auto)
--- a/src/ZF/Induct/ListN.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/ListN.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,19 +12,19 @@
   @{cite "paulin-tlca"}.
 \<close>
 
-consts listn :: "i=>i"
+consts listn :: "i\<Rightarrow>i"
 inductive
   domains "listn(A)" \<subseteq> "nat \<times> list(A)"
   intros
-    NilI: "<0,Nil> \<in> listn(A)"
-    ConsI: "\<lbrakk>a \<in> A; <n,l> \<in> listn(A)\<rbrakk> \<Longrightarrow> <succ(n), Cons(a,l)> \<in> listn(A)"
+    NilI: "\<langle>0,Nil\<rangle> \<in> listn(A)"
+    ConsI: "\<lbrakk>a \<in> A; \<langle>n,l\<rangle> \<in> listn(A)\<rbrakk> \<Longrightarrow> <succ(n), Cons(a,l)> \<in> listn(A)"
   type_intros nat_typechecks list.intros
 
 
 lemma list_into_listn: "l \<in> list(A) \<Longrightarrow> <length(l),l> \<in> listn(A)"
   by (induct set: list) (simp_all add: listn.intros)
 
-lemma listn_iff: "<n,l> \<in> listn(A) \<longleftrightarrow> l \<in> list(A) \<and> length(l)=n"
+lemma listn_iff: "\<langle>n,l\<rangle> \<in> listn(A) \<longleftrightarrow> l \<in> list(A) \<and> length(l)=n"
   apply (rule iffI)
    apply (erule listn.induct)
     apply auto
@@ -44,18 +44,18 @@
   done
 
 lemma listn_append:
-    "\<lbrakk><n,l> \<in> listn(A); <n',l'> \<in> listn(A)\<rbrakk> \<Longrightarrow> <n#+n', l@l'> \<in> listn(A)"
+    "\<lbrakk>\<langle>n,l\<rangle> \<in> listn(A); <n',l'> \<in> listn(A)\<rbrakk> \<Longrightarrow> <n#+n', l@l'> \<in> listn(A)"
   apply (erule listn.induct)
    apply (frule listn.dom_subset [THEN subsetD])
    apply (simp_all add: listn.intros)
   done
 
 inductive_cases
-      Nil_listn_case: "<i,Nil> \<in> listn(A)"
+      Nil_listn_case: "\<langle>i,Nil\<rangle> \<in> listn(A)"
   and Cons_listn_case: "<i,Cons(x,l)> \<in> listn(A)"
 
 inductive_cases
-      zero_listn_case: "<0,l> \<in> listn(A)"
+      zero_listn_case: "\<langle>0,l\<rangle> \<in> listn(A)"
   and succ_listn_case: "<succ(i),l> \<in> listn(A)"
 
 end
--- a/src/ZF/Induct/Multiset.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Multiset.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -13,68 +13,68 @@
 
 abbreviation (input)
   \<comment> \<open>Short cut for multiset space\<close>
-  Mult :: "i=>i" where
+  Mult :: "i\<Rightarrow>i" where
   "Mult(A) \<equiv> A -||> nat-{0}"
 
 definition
   (* This is the original "restrict" from ZF.thy.
      Restricts the function f to the domain A
      FIXME: adapt Multiset to the new "restrict". *)
-  funrestrict :: "[i,i] => i"  where
+  funrestrict :: "[i,i] \<Rightarrow> i"  where
   "funrestrict(f,A) \<equiv> \<lambda>x \<in> A. f`x"
 
 definition
   (* M is a multiset *)
-  multiset :: "i => o"  where
+  multiset :: "i \<Rightarrow> o"  where
   "multiset(M) \<equiv> \<exists>A. M \<in> A -> nat-{0} \<and> Finite(A)"
 
 definition
-  mset_of :: "i=>i"  where
+  mset_of :: "i\<Rightarrow>i"  where
   "mset_of(M) \<equiv> domain(M)"
 
 definition
-  munion    :: "[i, i] => i" (infixl \<open>+#\<close> 65)  where
+  munion    :: "[i, i] \<Rightarrow> i" (infixl \<open>+#\<close> 65)  where
   "M +# N \<equiv> \<lambda>x \<in> mset_of(M) \<union> mset_of(N).
      if x \<in> mset_of(M) \<inter> mset_of(N) then  (M`x) #+ (N`x)
      else (if x \<in> mset_of(M) then M`x else N`x)"
 
 definition
   (*convert a function to a multiset by eliminating 0*)
-  normalize :: "i => i"  where
+  normalize :: "i \<Rightarrow> i"  where
   "normalize(f) \<equiv>
        if (\<exists>A. f \<in> A -> nat \<and> Finite(A)) then
             funrestrict(f, {x \<in> mset_of(f). 0 < f`x})
        else 0"
 
 definition
-  mdiff  :: "[i, i] => i" (infixl \<open>-#\<close> 65)  where
+  mdiff  :: "[i, i] \<Rightarrow> i" (infixl \<open>-#\<close> 65)  where
   "M -# N \<equiv>  normalize(\<lambda>x \<in> mset_of(M).
                         if x \<in> mset_of(N) then M`x #- N`x else M`x)"
 
 definition
   (* set of elements of a multiset *)
-  msingle :: "i => i"    (\<open>{#_#}\<close>)  where
-  "{#a#} \<equiv> {<a, 1>}"
+  msingle :: "i \<Rightarrow> i"    (\<open>{#_#}\<close>)  where
+  "{#a#} \<equiv> {\<langle>a, 1\<rangle>}"
 
 definition
-  MCollect :: "[i, i=>o] => i"  (*comprehension*)  where
+  MCollect :: "[i, i\<Rightarrow>o] \<Rightarrow> i"  (*comprehension*)  where
   "MCollect(M, P) \<equiv> funrestrict(M, {x \<in> mset_of(M). P(x)})"
 
 definition
   (* Counts the number of occurrences of an element in a multiset *)
-  mcount :: "[i, i] => i"  where
+  mcount :: "[i, i] \<Rightarrow> i"  where
   "mcount(M, a) \<equiv> if a \<in> mset_of(M) then  M`a else 0"
 
 definition
-  msize :: "i => i"  where
-  "msize(M) \<equiv> setsum(%a. $# mcount(M,a), mset_of(M))"
+  msize :: "i \<Rightarrow> i"  where
+  "msize(M) \<equiv> setsum(\<lambda>a. $# mcount(M,a), mset_of(M))"
 
 abbreviation
-  melem :: "[i,i] => o"    (\<open>(_/ :# _)\<close> [50, 51] 50)  where
+  melem :: "[i,i] \<Rightarrow> o"    (\<open>(_/ :# _)\<close> [50, 51] 50)  where
   "a :# M \<equiv> a \<in> mset_of(M)"
 
 syntax
-  "_MColl" :: "[pttrn, i, o] => i" (\<open>(1{# _ \<in> _./ _#})\<close>)
+  "_MColl" :: "[pttrn, i, o] \<Rightarrow> i" (\<open>(1{# _ \<in> _./ _#})\<close>)
 translations
   "{#x \<in> M. P#}" == "CONST MCollect(M, \<lambda>x. P)"
 
@@ -83,28 +83,28 @@
 definition
    (* multirel1 has to be a set (not a predicate) so that we can form
       its transitive closure and reason about wf(.) and acc(.) *)
-  multirel1 :: "[i,i]=>i"  where
+  multirel1 :: "[i,i]\<Rightarrow>i"  where
   "multirel1(A, r) \<equiv>
-     {<M, N> \<in> Mult(A)*Mult(A).
+     {\<langle>M, N\<rangle> \<in> Mult(A)*Mult(A).
       \<exists>a \<in> A. \<exists>M0 \<in> Mult(A). \<exists>K \<in> Mult(A).
-      N=M0 +# {#a#} \<and> M=M0 +# K \<and> (\<forall>b \<in> mset_of(K). <b,a> \<in> r)}"
+      N=M0 +# {#a#} \<and> M=M0 +# K \<and> (\<forall>b \<in> mset_of(K). \<langle>b,a\<rangle> \<in> r)}"
 
 definition
-  multirel :: "[i, i] => i"  where
+  multirel :: "[i, i] \<Rightarrow> i"  where
   "multirel(A, r) \<equiv> multirel1(A, r)^+"
 
   (* ordinal multiset orderings *)
 
 definition
-  omultiset :: "i => o"  where
+  omultiset :: "i \<Rightarrow> o"  where
   "omultiset(M) \<equiv> \<exists>i. Ord(i) \<and> M \<in> Mult(field(Memrel(i)))"
 
 definition
-  mless :: "[i, i] => o" (infixl \<open><#\<close> 50)  where
-  "M <# N \<equiv>  \<exists>i. Ord(i) \<and> <M, N> \<in> multirel(field(Memrel(i)), Memrel(i))"
+  mless :: "[i, i] \<Rightarrow> o" (infixl \<open><#\<close> 50)  where
+  "M <# N \<equiv>  \<exists>i. Ord(i) \<and> \<langle>M, N\<rangle> \<in> multirel(field(Memrel(i)), Memrel(i))"
 
 definition
-  mle  :: "[i, i] => o"  (infixl \<open><#=\<close> 50)  where
+  mle  :: "[i, i] \<Rightarrow> o"  (infixl \<open><#=\<close> 50)  where
   "M <#= N \<equiv> (omultiset(M) \<and> M = N) | M <# N"
 
 
@@ -366,9 +366,9 @@
 apply (drule not_empty_multiset_imp_exist, assumption, clarify)
 apply (subgoal_tac "Finite (mset_of (M) - {a}) ")
  prefer 2 apply (simp add: Finite_Diff)
-apply (subgoal_tac "setsum (%x. $# mcount (M, x), cons (a, mset_of (M) -{a}))=#0")
+apply (subgoal_tac "setsum (\<lambda>x. $# mcount (M, x), cons (a, mset_of (M) -{a}))=#0")
  prefer 2 apply (simp add: cons_Diff, simp)
-apply (subgoal_tac "#0 $\<le> setsum (%x. $# mcount (M, x), mset_of (M) - {a}) ")
+apply (subgoal_tac "#0 $\<le> setsum (\<lambda>x. $# mcount (M, x), mset_of (M) - {a}) ")
 apply (rule_tac [2] g_zpos_imp_setsum_zpos)
 apply (auto simp add: Finite_Diff not_zless_iff_zle [THEN iff_sym] znegative_iff_zless_0 [THEN iff_sym])
 apply (rule not_zneg_int_of [THEN bexE])
@@ -376,8 +376,8 @@
 done
 
 lemma setsum_mcount_Int:
-     "Finite(A) \<Longrightarrow> setsum(%a. $# mcount(N, a), A \<inter> mset_of(N))
-                  = setsum(%a. $# mcount(N, a), A)"
+     "Finite(A) \<Longrightarrow> setsum(\<lambda>a. $# mcount(N, a), A \<inter> mset_of(N))
+                  = setsum(\<lambda>a. $# mcount(N, a), A)"
 apply (induct rule: Finite_induct)
  apply auto
 apply (subgoal_tac "Finite (B \<inter> mset_of (N))")
@@ -477,9 +477,9 @@
 lemma setsum_decr:
 "Finite(A)
   \<Longrightarrow>  (\<forall>M. multiset(M) \<longrightarrow>
-  (\<forall>a \<in> mset_of(M). setsum(%z. $# mcount(M(a:=M`a #- 1), z), A) =
-  (if a \<in> A then setsum(%z. $# mcount(M, z), A) $- #1
-           else setsum(%z. $# mcount(M, z), A))))"
+  (\<forall>a \<in> mset_of(M). setsum(\<lambda>z. $# mcount(M(a:=M`a #- 1), z), A) =
+  (if a \<in> A then setsum(\<lambda>z. $# mcount(M, z), A) $- #1
+           else setsum(\<lambda>z. $# mcount(M, z), A))))"
 apply (unfold multiset_def)
 apply (erule Finite_induct)
 apply (auto simp add: multiset_fun_iff)
@@ -493,19 +493,19 @@
 lemma setsum_decr2:
      "Finite(A)
       \<Longrightarrow> \<forall>M. multiset(M) \<longrightarrow> (\<forall>a \<in> mset_of(M).
-           setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) =
-           (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a
-            else setsum(%x. $# mcount(M, x), A)))"
+           setsum(\<lambda>x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A) =
+           (if a \<in> A then setsum(\<lambda>x. $# mcount(M, x), A) $- $# M`a
+            else setsum(\<lambda>x. $# mcount(M, x), A)))"
 apply (simp add: multiset_def)
 apply (erule Finite_induct)
 apply (auto simp add: multiset_fun_iff mcount_def mset_of_def)
 done
 
 lemma setsum_decr3: "\<lbrakk>Finite(A); multiset(M); a \<in> mset_of(M)\<rbrakk>
-      \<Longrightarrow> setsum(%x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) =
-          (if a \<in> A then setsum(%x. $# mcount(M, x), A) $- $# M`a
-           else setsum(%x. $# mcount(M, x), A))"
-apply (subgoal_tac "setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (%x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ")
+      \<Longrightarrow> setsum(\<lambda>x. $# mcount(funrestrict(M, mset_of(M)-{a}), x), A - {a}) =
+          (if a \<in> A then setsum(\<lambda>x. $# mcount(M, x), A) $- $# M`a
+           else setsum(\<lambda>x. $# mcount(M, x), A))"
+apply (subgoal_tac "setsum (\<lambda>x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A-{a}) = setsum (\<lambda>x. $# mcount (funrestrict (M, mset_of (M) -{a}),x),A) ")
 apply (rule_tac [2] setsum_Diff [symmetric])
 apply (rule sym, rule ssubst, blast)
 apply (rule sym, drule setsum_decr2, auto)
@@ -530,22 +530,22 @@
 done
 
 lemma multiset_induct_aux:
-  assumes prem1: "\<And>M a. \<lbrakk>multiset(M); a\<notin>mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(cons(<a, 1>, M))"
+  assumes prem1: "\<And>M a. \<lbrakk>multiset(M); a\<notin>mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(cons(\<langle>a, 1\<rangle>, M))"
       and prem2: "\<And>M b. \<lbrakk>multiset(M); b \<in> mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(M(b:= M`b #+ 1))"
   shows
   "\<lbrakk>n \<in> nat; P(0)\<rbrakk>
      \<Longrightarrow> (\<forall>M. multiset(M)\<longrightarrow>
-  (setsum(%x. $# mcount(M, x), {x \<in> mset_of(M). 0 < M`x}) = $# n) \<longrightarrow> P(M))"
+  (setsum(\<lambda>x. $# mcount(M, x), {x \<in> mset_of(M). 0 < M`x}) = $# n) \<longrightarrow> P(M))"
 apply (erule nat_induct, clarify)
 apply (frule msize_eq_0_iff)
 apply (auto simp add: mset_of_def multiset_def multiset_fun_iff msize_def)
-apply (subgoal_tac "setsum (%x. $# mcount (M, x), A) =$# succ (x) ")
+apply (subgoal_tac "setsum (\<lambda>x. $# mcount (M, x), A) =$# succ (x) ")
 apply (drule setsum_succD, auto)
 apply (case_tac "1 <M`a")
 apply (drule_tac [2] not_lt_imp_le)
 apply (simp_all add: nat_le_1_cases)
 apply (subgoal_tac "M= (M (a:=M`a #- 1)) (a:= (M (a:=M`a #- 1))`a #+ 1) ")
-apply (rule_tac [2] A = A and B = "%x. nat" and D = "%x. nat" in fun_extension)
+apply (rule_tac [2] A = A and B = "\<lambda>x. nat" and D = "\<lambda>x. nat" in fun_extension)
 apply (rule_tac [3] update_type)+
 apply (simp_all (no_asm_simp))
  apply (rule_tac [2] impI)
@@ -581,7 +581,7 @@
  apply (subgoal_tac "cons (a, A-{a}) = A")
   apply force
   apply force
-apply (rule_tac a = "cons (<a, 1>, funrestrict (M, A - {a}))" in ssubst)
+apply (rule_tac a = "cons (\<langle>a, 1\<rangle>, funrestrict (M, A - {a}))" in ssubst)
 apply simp
 apply (frule multiset_funrestict, assumption)
 apply (rule prem1, assumption)
@@ -601,7 +601,7 @@
 
 lemma multiset_induct2:
   "\<lbrakk>multiset(M); P(0);
-    (\<And>M a. \<lbrakk>multiset(M); a\<notin>mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(cons(<a, 1>, M)));
+    (\<And>M a. \<lbrakk>multiset(M); a\<notin>mset_of(M); P(M)\<rbrakk> \<Longrightarrow> P(cons(\<langle>a, 1\<rangle>, M)));
     (\<And>M b. \<lbrakk>multiset(M); b \<in> mset_of(M);  P(M)\<rbrakk> \<Longrightarrow> P(M(b:= M`b #+ 1)))\<rbrakk>
      \<Longrightarrow> P(M)"
 apply (subgoal_tac "\<exists>n \<in> nat. setsum (\<lambda>x. $# mcount (M, x), {x \<in> mset_of (M) . 0 < M ` x}) = $# n")
@@ -614,7 +614,7 @@
 done
 
 lemma munion_single_case1:
-     "\<lbrakk>multiset(M); a \<notin>mset_of(M)\<rbrakk> \<Longrightarrow> M +# {#a#} = cons(<a, 1>, M)"
+     "\<lbrakk>multiset(M); a \<notin>mset_of(M)\<rbrakk> \<Longrightarrow> M +# {#a#} = cons(\<langle>a, 1\<rangle>, M)"
 apply (simp add: multiset_def msingle_def)
 apply (auto simp add: munion_def)
 apply (unfold mset_of_def, simp)
@@ -726,10 +726,10 @@
 by (auto simp add: multirel1_def)
 
 lemma multirel1_iff:
-" <N, M> \<in> multirel1(A, r) \<longleftrightarrow>
+" \<langle>N, M\<rangle> \<in> multirel1(A, r) \<longleftrightarrow>
   (\<exists>a. a \<in> A \<and>
   (\<exists>M0. M0 \<in> Mult(A) \<and> (\<exists>K. K \<in> Mult(A) \<and>
-   M=M0 +# {#a#} \<and> N=M0 +# K \<and> (\<forall>b \<in> mset_of(K). <b,a> \<in> r))))"
+   M=M0 +# {#a#} \<and> N=M0 +# K \<and> (\<forall>b \<in> mset_of(K). \<langle>b,a\<rangle> \<in> r))))"
 by (auto simp add: multirel1_def Mult_iff_multiset Bex_def)
 
 
@@ -762,12 +762,12 @@
 
 subsection\<open>Toward the proof of well-foundedness of multirel1\<close>
 
-lemma not_less_0 [iff]: "<M,0> \<notin> multirel1(A, r)"
+lemma not_less_0 [iff]: "\<langle>M,0\<rangle> \<notin> multirel1(A, r)"
 by (auto simp add: multirel1_def Mult_iff_multiset)
 
 lemma less_munion: "\<lbrakk><N, M0 +# {#a#}> \<in> multirel1(A, r); M0 \<in> Mult(A)\<rbrakk> \<Longrightarrow>
-  (\<exists>M. <M, M0> \<in> multirel1(A, r) \<and> N = M +# {#a#}) |
-  (\<exists>K. K \<in> Mult(A) \<and> (\<forall>b \<in> mset_of(K). <b, a> \<in> r) \<and> N = M0 +# K)"
+  (\<exists>M. \<langle>M, M0\<rangle> \<in> multirel1(A, r) \<and> N = M +# {#a#}) |
+  (\<exists>K. K \<in> Mult(A) \<and> (\<forall>b \<in> mset_of(K). \<langle>b, a\<rangle> \<in> r) \<and> N = M0 +# K)"
 apply (frule multirel1_type [THEN subsetD])
 apply (simp add: multirel1_iff)
 apply (auto simp add: munion_eq_conv_exist)
@@ -787,10 +787,10 @@
 lemma acc_0: "acc(0)=0"
 by (auto intro!: equalityI dest: acc.dom_subset [THEN subsetD])
 
-lemma lemma1: "\<lbrakk>\<forall>b \<in> A. <b,a> \<in> r \<longrightarrow>
+lemma lemma1: "\<lbrakk>\<forall>b \<in> A. \<langle>b,a\<rangle> \<in> r \<longrightarrow>
     (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#}:acc(multirel1(A, r)));
     M0 \<in> acc(multirel1(A, r)); a \<in> A;
-    \<forall>M. <M,M0> \<in> multirel1(A, r) \<longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r))\<rbrakk>
+    \<forall>M. \<langle>M,M0\<rangle> \<in> multirel1(A, r) \<longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r))\<rbrakk>
   \<Longrightarrow> M0 +# {#a#} \<in> acc(multirel1(A, r))"
 apply (subgoal_tac "M0 \<in> Mult(A) ")
  prefer 2
@@ -801,7 +801,7 @@
 apply (rename_tac "N")
 apply (drule less_munion, blast)
 apply (auto simp add: Mult_iff_multiset)
-apply (erule_tac P = "\<forall>x \<in> mset_of (K) . <x, a> \<in> r" in rev_mp)
+apply (erule_tac P = "\<forall>x \<in> mset_of (K) . \<langle>x, a\<rangle> \<in> r" in rev_mp)
 apply (erule_tac P = "mset_of (K) \<subseteq>A" in rev_mp)
 apply (erule_tac M = K in multiset_induct)
 (* three subgoals *)
@@ -813,13 +813,13 @@
 apply (subgoal_tac "aa \<in> A")
 prefer 2 apply blast
 apply (drule_tac x = "M0 +# M" and P =
-       "%x. x \<in> acc(multirel1(A, r)) \<longrightarrow> Q(x)" for Q in spec)
+       "\<lambda>x. x \<in> acc(multirel1(A, r)) \<longrightarrow> Q(x)" for Q in spec)
 apply (simp add: munion_assoc [symmetric])
 (* subgoal 3 \<in> additional conditions *)
 apply (auto intro!: multirel1_base [THEN fieldI2] simp add: Mult_iff_multiset)
 done
 
-lemma lemma2: "\<lbrakk>\<forall>b \<in> A. <b,a> \<in> r
+lemma lemma2: "\<lbrakk>\<forall>b \<in> A. \<langle>b,a\<rangle> \<in> r
    \<longrightarrow> (\<forall>M \<in> acc(multirel1(A, r)). M +# {#b#} :acc(multirel1(A, r)));
         M \<in> acc(multirel1(A, r)); a \<in> A\<rbrakk> \<Longrightarrow> M +# {#a#} \<in> acc(multirel1(A, r))"
 apply (erule acc_induct)
@@ -914,12 +914,12 @@
 (* One direction *)
 
 lemma multirel_implies_one_step:
-"<M,N> \<in> multirel(A, r) \<Longrightarrow>
+"\<langle>M,N\<rangle> \<in> multirel(A, r) \<Longrightarrow>
      trans[A](r) \<longrightarrow>
      (\<exists>I J K.
          I \<in> Mult(A) \<and> J \<in> Mult(A) \<and>  K \<in> Mult(A) \<and>
          N = I +# J \<and> M = I +# K \<and> J \<noteq> 0 \<and>
-        (\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r))"
+        (\<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). \<langle>k,j\<rangle> \<in> r))"
 apply (simp add: multirel_def Ball_def Bex_def)
 apply (erule converse_trancl_induct)
 apply (simp_all add: multirel1_iff Mult_iff_multiset)
@@ -936,11 +936,11 @@
 apply (rule_tac x = " (Ka -# {#a#}) +# K" in exI, simp (no_asm_simp))
 apply (simp_all add: Un_subset_iff)
 apply (simp (no_asm_simp) add: munion_assoc [symmetric])
-apply (drule_tac t = "%M. M-#{#a#}" in subst_context)
+apply (drule_tac t = "\<lambda>M. M-#{#a#}" in subst_context)
 apply (simp add: mdiff_union_single_conv melem_diff_single, clarify)
 apply (erule disjE, simp)
 apply (erule disjE, simp)
-apply (drule_tac x = a and P = "%x. x :# Ka \<longrightarrow> Q(x)" for Q in spec)
+apply (drule_tac x = a and P = "\<lambda>x. x :# Ka \<longrightarrow> Q(x)" for Q in spec)
 apply clarify
 apply (rule_tac x = xa in exI)
 apply (simp (no_asm_simp))
@@ -955,7 +955,7 @@
 apply (rule conjI)
 apply (simp (no_asm_simp) add: multiset_equality mcount_elem [THEN succ_pred_eq_self])
 apply (rule conjI)
-apply (drule_tac t = "%M. M-#{#a#}" in subst_context)
+apply (drule_tac t = "\<lambda>M. M-#{#a#}" in subst_context)
 apply (simp add: mdiff_union_inverse2)
 apply (simp_all (no_asm_simp) add: multiset_equality)
 apply (rule diff_add_commute [symmetric])
@@ -984,7 +984,7 @@
 "n \<in> nat \<Longrightarrow>
    (\<forall>I J K.
     I \<in> Mult(A) \<and> J \<in> Mult(A) \<and> K \<in> Mult(A) \<and>
-   (msize(J) = $# n \<and> J \<noteq>0 \<and>  (\<forall>k \<in> mset_of(K).  \<exists>j \<in> mset_of(J). <k, j> \<in> r))
+   (msize(J) = $# n \<and> J \<noteq>0 \<and>  (\<forall>k \<in> mset_of(K).  \<exists>j \<in> mset_of(J). \<langle>k, j\<rangle> \<in> r))
     \<longrightarrow> <I +# K, I +# J> \<in> multirel(A, r))"
 apply (simp add: Mult_iff_multiset)
 apply (erule nat_induct, clarify)
@@ -1002,19 +1002,19 @@
 (*Now we know J' \<noteq>  0*)
 apply (drule sym, rotate_tac -1, simp)
 apply (erule_tac V = "$# x = msize (J') " in thin_rl)
-apply (frule_tac M = K and P = "%x. <x,a> \<in> r" in multiset_partition)
+apply (frule_tac M = K and P = "\<lambda>x. \<langle>x,a\<rangle> \<in> r" in multiset_partition)
 apply (erule_tac P = "\<forall>k \<in> mset_of (K) . P(k)" for P in rev_mp)
 apply (erule ssubst)
 apply (simp add: Ball_def, auto)
-apply (subgoal_tac "< (I +# {# x \<in> K. <x, a> \<in> r#}) +# {# x \<in> K. <x, a> \<notin> r#}, (I +# {# x \<in> K. <x, a> \<in> r#}) +# J'> \<in> multirel(A, r) ")
+apply (subgoal_tac "< (I +# {# x \<in> K. \<langle>x, a\<rangle> \<in> r#}) +# {# x \<in> K. \<langle>x, a\<rangle> \<notin> r#}, (I +# {# x \<in> K. \<langle>x, a\<rangle> \<in> r#}) +# J'> \<in> multirel(A, r) ")
  prefer 2
- apply (drule_tac x = "I +# {# x \<in> K. <x, a> \<in> r#}" in spec)
+ apply (drule_tac x = "I +# {# x \<in> K. \<langle>x, a\<rangle> \<in> r#}" in spec)
  apply (rotate_tac -1)
  apply (drule_tac x = "J'" in spec)
  apply (rotate_tac -1)
- apply (drule_tac x = "{# x \<in> K. <x, a> \<notin> r#}" in spec, simp) apply blast
+ apply (drule_tac x = "{# x \<in> K. \<langle>x, a\<rangle> \<notin> r#}" in spec, simp) apply blast
 apply (simp add: munion_assoc [symmetric] multirel_def)
-apply (rule_tac b = "I +# {# x \<in> K. <x, a> \<in> r#} +# J'" in trancl_trans, blast)
+apply (rule_tac b = "I +# {# x \<in> K. \<langle>x, a\<rangle> \<in> r#} +# J'" in trancl_trans, blast)
 apply (rule r_into_trancl)
 apply (simp add: multirel1_iff Mult_iff_multiset)
 apply (rule_tac x = a in exI)
@@ -1024,7 +1024,7 @@
 done
 
 lemma one_step_implies_multirel:
-     "\<lbrakk>J \<noteq> 0;  \<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). <k,j> \<in> r;
+     "\<lbrakk>J \<noteq> 0;  \<forall>k \<in> mset_of(K). \<exists>j \<in> mset_of(J). \<langle>k,j\<rangle> \<in> r;
          I \<in> Mult(A); J \<in> Mult(A); K \<in> Mult(A)\<rbrakk>
       \<Longrightarrow> <I+#K, I+#J> \<in> multirel(A, r)"
 apply (subgoal_tac "multiset (J) ")
@@ -1038,7 +1038,7 @@
 (*irreflexivity*)
 
 lemma multirel_irrefl_lemma:
-     "Finite(A) \<Longrightarrow> part_ord(A, r) \<longrightarrow> (\<forall>x \<in> A. \<exists>y \<in> A. <x,y> \<in> r) \<longrightarrow>A=0"
+     "Finite(A) \<Longrightarrow> part_ord(A, r) \<longrightarrow> (\<forall>x \<in> A. \<exists>y \<in> A. \<langle>x,y\<rangle> \<in> r) \<longrightarrow>A=0"
 apply (erule Finite_induct)
 apply (auto dest: subset_consI [THEN [2] part_ord_subset])
 apply (auto simp add: part_ord_def irrefl_def)
@@ -1066,7 +1066,7 @@
 done
 
 lemma multirel_trans:
- "\<lbrakk><M, N> \<in> multirel(A, r); <N, K> \<in> multirel(A, r)\<rbrakk> \<Longrightarrow>  <M, K> \<in> multirel(A,r)"
+ "\<lbrakk>\<langle>M, N\<rangle> \<in> multirel(A, r); \<langle>N, K\<rangle> \<in> multirel(A, r)\<rbrakk> \<Longrightarrow>  \<langle>M, K\<rangle> \<in> multirel(A,r)"
 apply (simp add: multirel_def)
 apply (blast intro: trancl_trans)
 done
@@ -1084,7 +1084,7 @@
 (** Monotonicity of multiset union **)
 
 lemma munion_multirel1_mono:
-"\<lbrakk><M,N> \<in> multirel1(A, r); K \<in> Mult(A)\<rbrakk> \<Longrightarrow> <K +# M, K +# N> \<in> multirel1(A, r)"
+"\<lbrakk>\<langle>M,N\<rangle> \<in> multirel1(A, r); K \<in> Mult(A)\<rbrakk> \<Longrightarrow> <K +# M, K +# N> \<in> multirel1(A, r)"
 apply (frule multirel1_type [THEN subsetD])
 apply (auto simp add: multirel1_iff Mult_iff_multiset)
 apply (rule_tac x = a in exI)
@@ -1096,11 +1096,11 @@
 done
 
 lemma munion_multirel_mono2:
- "\<lbrakk><M, N> \<in> multirel(A, r); K \<in> Mult(A)\<rbrakk>\<Longrightarrow><K +# M, K +# N> \<in> multirel(A, r)"
+ "\<lbrakk>\<langle>M, N\<rangle> \<in> multirel(A, r); K \<in> Mult(A)\<rbrakk>\<Longrightarrow><K +# M, K +# N> \<in> multirel(A, r)"
 apply (frule multirel_type [THEN subsetD])
 apply (simp (no_asm_use) add: multirel_def)
 apply clarify
-apply (drule_tac psi = "<M,N> \<in> multirel1 (A, r) ^+" in asm_rl)
+apply (drule_tac psi = "\<langle>M,N\<rangle> \<in> multirel1 (A, r) ^+" in asm_rl)
 apply (erule rev_mp)
 apply (erule rev_mp)
 apply (erule rev_mp)
@@ -1115,16 +1115,16 @@
 done
 
 lemma munion_multirel_mono1:
-     "\<lbrakk><M, N> \<in> multirel(A, r); K \<in> Mult(A)\<rbrakk> \<Longrightarrow> <M +# K, N +# K> \<in> multirel(A, r)"
+     "\<lbrakk>\<langle>M, N\<rangle> \<in> multirel(A, r); K \<in> Mult(A)\<rbrakk> \<Longrightarrow> <M +# K, N +# K> \<in> multirel(A, r)"
 apply (frule multirel_type [THEN subsetD])
-apply (rule_tac P = "%x. <x,u> \<in> multirel(A, r)" for u in munion_commute [THEN subst])
+apply (rule_tac P = "\<lambda>x. \<langle>x,u\<rangle> \<in> multirel(A, r)" for u in munion_commute [THEN subst])
 apply (subst munion_commute [of N])
 apply (rule munion_multirel_mono2)
 apply (auto simp add: Mult_iff_multiset)
 done
 
 lemma munion_multirel_mono:
-     "\<lbrakk><M,K> \<in> multirel(A, r); <N,L> \<in> multirel(A, r)\<rbrakk>
+     "\<lbrakk>\<langle>M,K\<rangle> \<in> multirel(A, r); \<langle>N,L\<rangle> \<in> multirel(A, r)\<rbrakk>
       \<Longrightarrow> <M +# N, K +# L> \<in> multirel(A, r)"
 apply (subgoal_tac "M \<in> Mult(A) \<and> N \<in> Mult(A) \<and> K \<in> Mult(A) \<and> L \<in> Mult(A) ")
 prefer 2 apply (blast dest: multirel_type [THEN subsetD])
--- a/src/ZF/Induct/Mutil.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Mutil.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -14,13 +14,13 @@
 
 consts
   domino :: i
-  tiling :: "i => i"
+  tiling :: "i \<Rightarrow> i"
 
 inductive
   domains "domino" \<subseteq> "Pow(nat \<times> nat)"
   intros
-    horiz: "\<lbrakk>i \<in> nat;  j \<in> nat\<rbrakk> \<Longrightarrow> {<i,j>, <i,succ(j)>} \<in> domino"
-    vertl: "\<lbrakk>i \<in> nat;  j \<in> nat\<rbrakk> \<Longrightarrow> {<i,j>, <succ(i),j>} \<in> domino"
+    horiz: "\<lbrakk>i \<in> nat;  j \<in> nat\<rbrakk> \<Longrightarrow> {\<langle>i,j\<rangle>, <i,succ(j)>} \<in> domino"
+    vertl: "\<lbrakk>i \<in> nat;  j \<in> nat\<rbrakk> \<Longrightarrow> {\<langle>i,j\<rangle>, <succ(i),j>} \<in> domino"
   type_intros empty_subsetI cons_subsetI PowI SigmaI nat_succI
 
 inductive
@@ -32,13 +32,13 @@
   type_elims PowD [elim_format]
 
 definition
-  evnodd :: "[i, i] => i"  where
-  "evnodd(A,b) \<equiv> {z \<in> A. \<exists>i j. z = <i,j> \<and> (i #+ j) mod 2 = b}"
+  evnodd :: "[i, i] \<Rightarrow> i"  where
+  "evnodd(A,b) \<equiv> {z \<in> A. \<exists>i j. z = \<langle>i,j\<rangle> \<and> (i #+ j) mod 2 = b}"
 
 
 subsection \<open>Basic properties of evnodd\<close>
 
-lemma evnodd_iff: "<i,j>: evnodd(A,b) \<longleftrightarrow> <i,j>: A \<and> (i#+j) mod 2 = b"
+lemma evnodd_iff: "\<langle>i,j\<rangle>: evnodd(A,b) \<longleftrightarrow> \<langle>i,j\<rangle>: A \<and> (i#+j) mod 2 = b"
   by (unfold evnodd_def) blast
 
 lemma evnodd_subset: "evnodd(A, b) \<subseteq> A"
@@ -54,8 +54,8 @@
   by (simp add: evnodd_def Collect_Diff)
 
 lemma evnodd_cons [simp]:
-  "evnodd(cons(<i,j>,C), b) =
-    (if (i#+j) mod 2 = b then cons(<i,j>, evnodd(C,b)) else evnodd(C,b))"
+  "evnodd(cons(\<langle>i,j\<rangle>,C), b) =
+    (if (i#+j) mod 2 = b then cons(\<langle>i,j\<rangle>, evnodd(C,b)) else evnodd(C,b))"
   by (simp add: evnodd_def Collect_cons)
 
 lemma evnodd_0 [simp]: "evnodd(0, b) = 0"
@@ -143,7 +143,7 @@
 
 theorem mutil_not_tiling: "\<lbrakk>m \<in> nat;  n \<in> nat;
          t = (succ(m)#+succ(m))*(succ(n)#+succ(n));
-         t' = t - {<0,0>} - {<succ(m#+m), succ(n#+n)>}\<rbrakk>
+         t' = t - {\<langle>0,0\<rangle>} - {<succ(m#+m), succ(n#+n)>}\<rbrakk>
       \<Longrightarrow> t' \<notin> tiling(domino)"
   apply (rule notI)
   apply (drule tiling_domino_0_1)
--- a/src/ZF/Induct/Ntree.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Ntree.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -13,9 +13,9 @@
 \<close>
 
 consts
-  ntree :: "i => i"
-  maptree :: "i => i"
-  maptree2 :: "[i, i] => i"
+  ntree :: "i \<Rightarrow> i"
+  maptree :: "i \<Rightarrow> i"
+  maptree2 :: "[i, i] \<Rightarrow> i"
 
 datatype "ntree(A)" = Branch ("a \<in> A", "h \<in> (\<Union>n \<in> nat. n -> ntree(A))")
   monos UN_mono [OF subset_refl Pi_mono]  \<comment> \<open>MUST have this form\<close>
@@ -31,12 +31,12 @@
   type_intros FiniteFun_in_univ'
 
 definition
-  ntree_rec :: "[[i, i, i] => i, i] => i"  where
+  ntree_rec :: "[[i, i, i] \<Rightarrow> i, i] \<Rightarrow> i"  where
   "ntree_rec(b) \<equiv>
     Vrecursor(\<lambda>pr. ntree_case(\<lambda>x h. b(x, h, \<lambda>i \<in> domain(h). pr`(h`i))))"
 
 definition
-  ntree_copy :: "i => i"  where
+  ntree_copy :: "i \<Rightarrow> i"  where
   "ntree_copy(z) \<equiv> ntree_rec(\<lambda>x h r. Branch(x,r), z)"
 
 
--- a/src/ZF/Induct/Primrec.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Primrec.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -21,32 +21,32 @@
   "SC \<equiv> \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. succ(x), l)"
 
 definition
-  CONSTANT :: "i=>i"  where
+  CONSTANT :: "i\<Rightarrow>i"  where
   "CONSTANT(k) \<equiv> \<lambda>l \<in> list(nat). k"
 
 definition
-  PROJ :: "i=>i"  where
+  PROJ :: "i\<Rightarrow>i"  where
   "PROJ(i) \<equiv> \<lambda>l \<in> list(nat). list_case(0, \<lambda>x xs. x, drop(i,l))"
 
 definition
-  COMP :: "[i,i]=>i"  where
+  COMP :: "[i,i]\<Rightarrow>i"  where
   "COMP(g,fs) \<equiv> \<lambda>l \<in> list(nat). g ` map(\<lambda>f. f`l, fs)"
 
 definition
-  PREC :: "[i,i]=>i"  where
+  PREC :: "[i,i]\<Rightarrow>i"  where
   "PREC(f,g) \<equiv>
      \<lambda>l \<in> list(nat). list_case(0,
                       \<lambda>x xs. rec(x, f`xs, \<lambda>y r. g ` Cons(r, Cons(y, xs))), l)"
   \<comment> \<open>Note that \<open>g\<close> is applied first to \<^term>\<open>PREC(f,g)`y\<close> and then to \<open>y\<close>!\<close>
 
 consts
-  ACK :: "i=>i"
+  ACK :: "i\<Rightarrow>i"
 primrec
   "ACK(0) = SC"
   "ACK(succ(i)) = PREC (CONSTANT (ACK(i) ` [1]), COMP(ACK(i), [PROJ(0)]))"
 
 abbreviation
-  ack :: "[i,i]=>i" where
+  ack :: "[i,i]\<Rightarrow>i" where
   "ack(x,y) \<equiv> ACK(x) ` [y]"
 
 
--- a/src/ZF/Induct/PropLog.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/PropLog.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -27,26 +27,26 @@
 datatype propn =
     Fls
   | Var ("n \<in> nat")    (\<open>#_\<close> [100] 100)
-  | Imp ("p \<in> propn", "q \<in> propn")    (infixr \<open>=>\<close> 90)
+  | Imp ("p \<in> propn", "q \<in> propn")    (infixr \<open>\<Rightarrow>\<close> 90)
 
 
 subsection \<open>The proof system\<close>
 
-consts thms     :: "i => i"
+consts thms     :: "i \<Rightarrow> i"
 
 abbreviation
-  thms_syntax :: "[i,i] => o"    (infixl \<open>|-\<close> 50)
+  thms_syntax :: "[i,i] \<Rightarrow> o"    (infixl \<open>|-\<close> 50)
   where "H |- p \<equiv> p \<in> thms(H)"
 
 inductive
   domains "thms(H)" \<subseteq> "propn"
   intros
     H:  "\<lbrakk>p \<in> H;  p \<in> propn\<rbrakk> \<Longrightarrow> H |- p"
-    K:  "\<lbrakk>p \<in> propn;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- p=>q=>p"
+    K:  "\<lbrakk>p \<in> propn;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- p\<Rightarrow>q\<Rightarrow>p"
     S:  "\<lbrakk>p \<in> propn;  q \<in> propn;  r \<in> propn\<rbrakk>
-         \<Longrightarrow> H |- (p=>q=>r) => (p=>q) => p=>r"
-    DN: "p \<in> propn \<Longrightarrow> H |- ((p=>Fls) => Fls) => p"
-    MP: "\<lbrakk>H |- p=>q;  H |- p;  p \<in> propn;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
+         \<Longrightarrow> H |- (p\<Rightarrow>q\<Rightarrow>r) \<Rightarrow> (p\<Rightarrow>q) \<Rightarrow> p\<Rightarrow>r"
+    DN: "p \<in> propn \<Longrightarrow> H |- ((p\<Rightarrow>Fls) \<Rightarrow> Fls) \<Rightarrow> p"
+    MP: "\<lbrakk>H |- p\<Rightarrow>q;  H |- p;  p \<in> propn;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
   type_intros "propn.intros"
 
 declare propn.intros [simp]
@@ -57,14 +57,14 @@
 subsubsection \<open>Semantics of propositional logic.\<close>
 
 consts
-  is_true_fun :: "[i,i] => i"
+  is_true_fun :: "[i,i] \<Rightarrow> i"
 primrec
   "is_true_fun(Fls, t) = 0"
   "is_true_fun(Var(v), t) = (if v \<in> t then 1 else 0)"
-  "is_true_fun(p=>q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)"
+  "is_true_fun(p\<Rightarrow>q, t) = (if is_true_fun(p,t) = 1 then is_true_fun(q,t) else 1)"
 
 definition
-  is_true :: "[i,i] => o"  where
+  is_true :: "[i,i] \<Rightarrow> o"  where
   "is_true(p,t) \<equiv> is_true_fun(p,t) = 1"
   \<comment> \<open>this definition is required since predicates can't be recursive\<close>
 
@@ -74,7 +74,7 @@
 lemma is_true_Var [simp]: "is_true(#v,t) \<longleftrightarrow> v \<in> t"
   by (simp add: is_true_def)
 
-lemma is_true_Imp [simp]: "is_true(p=>q,t) \<longleftrightarrow> (is_true(p,t)\<longrightarrow>is_true(q,t))"
+lemma is_true_Imp [simp]: "is_true(p\<Rightarrow>q,t) \<longleftrightarrow> (is_true(p,t)\<longrightarrow>is_true(q,t))"
   by (simp add: is_true_def)
 
 
@@ -86,7 +86,7 @@
 \<close>
 
 definition
-  logcon :: "[i,i] => o"    (infixl \<open>|=\<close> 50)  where
+  logcon :: "[i,i] \<Rightarrow> o"    (infixl \<open>|=\<close> 50)  where
   "H |= p \<equiv> \<forall>t. (\<forall>q \<in> H. is_true(q,t)) \<longrightarrow> is_true(p,t)"
 
 
@@ -96,11 +96,11 @@
 \<close>
 
 consts
-  hyps :: "[i,i] => i"
+  hyps :: "[i,i] \<Rightarrow> i"
 primrec
   "hyps(Fls, t) = 0"
-  "hyps(Var(v), t) = (if v \<in> t then {#v} else {#v=>Fls})"
-  "hyps(p=>q, t) = hyps(p,t) \<union> hyps(q,t)"
+  "hyps(Var(v), t) = (if v \<in> t then {#v} else {#v\<Rightarrow>Fls})"
+  "hyps(p\<Rightarrow>q, t) = hyps(p,t) \<union> hyps(q,t)"
 
 
 
@@ -115,15 +115,15 @@
 
 lemmas thms_in_pl = thms.dom_subset [THEN subsetD]
 
-inductive_cases ImpE: "p=>q \<in> propn"
+inductive_cases ImpE: "p\<Rightarrow>q \<in> propn"
 
-lemma thms_MP: "\<lbrakk>H |- p=>q;  H |- p\<rbrakk> \<Longrightarrow> H |- q"
+lemma thms_MP: "\<lbrakk>H |- p\<Rightarrow>q;  H |- p\<rbrakk> \<Longrightarrow> H |- q"
   \<comment> \<open>Stronger Modus Ponens rule: no typechecking!\<close>
   apply (rule thms.MP)
      apply (erule asm_rl thms_in_pl thms_in_pl [THEN ImpE])+
   done
 
-lemma thms_I: "p \<in> propn \<Longrightarrow> H |- p=>p"
+lemma thms_I: "p \<in> propn \<Longrightarrow> H |- p\<Rightarrow>p"
   \<comment> \<open>Rule is called \<open>I\<close> for Identity Combinator, not for Introduction.\<close>
   apply (rule thms.S [THEN thms_MP, THEN thms_MP])
       apply (rule_tac [5] thms.K)
@@ -144,13 +144,13 @@
 lemmas weaken_left_Un1  = Un_upper1 [THEN weaken_left]
 lemmas weaken_left_Un2  = Un_upper2 [THEN weaken_left]
 
-lemma weaken_right: "\<lbrakk>H |- q;  p \<in> propn\<rbrakk> \<Longrightarrow> H |- p=>q"
+lemma weaken_right: "\<lbrakk>H |- q;  p \<in> propn\<rbrakk> \<Longrightarrow> H |- p\<Rightarrow>q"
   by (simp_all add: thms.K [THEN thms_MP] thms_in_pl)
 
 
 subsubsection \<open>The deduction theorem\<close>
 
-theorem deduction: "\<lbrakk>cons(p,H) |- q;  p \<in> propn\<rbrakk> \<Longrightarrow>  H |- p=>q"
+theorem deduction: "\<lbrakk>cons(p,H) |- q;  p \<in> propn\<rbrakk> \<Longrightarrow>  H |- p\<Rightarrow>q"
   apply (erule thms.induct)
       apply (blast intro: thms_I thms.H [THEN weaken_right])
      apply (blast intro: thms.K [THEN weaken_right])
@@ -173,7 +173,7 @@
     apply (simp_all add: propn.intros)
   done
 
-lemma thms_notE: "\<lbrakk>H |- p=>Fls;  H |- p;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
+lemma thms_notE: "\<lbrakk>H |- p\<Rightarrow>Fls;  H |- p;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
   by (erule thms_MP [THEN thms_FlsE])
 
 
@@ -190,16 +190,16 @@
 
 subsubsection \<open>Towards the completeness proof\<close>
 
-lemma Fls_Imp: "\<lbrakk>H |- p=>Fls; q \<in> propn\<rbrakk> \<Longrightarrow> H |- p=>q"
+lemma Fls_Imp: "\<lbrakk>H |- p\<Rightarrow>Fls; q \<in> propn\<rbrakk> \<Longrightarrow> H |- p\<Rightarrow>q"
   apply (frule thms_in_pl)
   apply (rule deduction)
    apply (rule weaken_left_cons [THEN thms_notE])
      apply (blast intro: thms.H elim: ImpE)+
   done
 
-lemma Imp_Fls: "\<lbrakk>H |- p;  H |- q=>Fls\<rbrakk> \<Longrightarrow> H |- (p=>q)=>Fls"
+lemma Imp_Fls: "\<lbrakk>H |- p;  H |- q\<Rightarrow>Fls\<rbrakk> \<Longrightarrow> H |- (p\<Rightarrow>q)\<Rightarrow>Fls"
   apply (frule thms_in_pl)
-  apply (frule thms_in_pl [of concl: "q=>Fls"])
+  apply (frule thms_in_pl [of concl: "q\<Rightarrow>Fls"])
   apply (rule deduction)
    apply (erule weaken_left_cons [THEN thms_MP])
    apply (rule consI1 [THEN thms.H, THEN thms_MP])
@@ -207,7 +207,7 @@
   done
 
 lemma hyps_thms_if:
-    "p \<in> propn \<Longrightarrow> hyps(p,t) |- (if is_true(p,t) then p else p=>Fls)"
+    "p \<in> propn \<Longrightarrow> hyps(p,t) |- (if is_true(p,t) then p else p\<Rightarrow>Fls)"
   \<comment> \<open>Typical example of strengthening the induction statement.\<close>
   apply simp
   apply (induct_tac p)
@@ -234,14 +234,14 @@
 \<close>
 
 lemma thms_excluded_middle:
-    "\<lbrakk>p \<in> propn;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- (p=>q) => ((p=>Fls)=>q) => q"
+    "\<lbrakk>p \<in> propn;  q \<in> propn\<rbrakk> \<Longrightarrow> H |- (p\<Rightarrow>q) \<Rightarrow> ((p\<Rightarrow>Fls)\<Rightarrow>q) \<Rightarrow> q"
   apply (rule deduction [THEN deduction])
     apply (rule thms.DN [THEN thms_MP])
      apply (best intro!: propn_SIs intro: propn_Is)+
   done
 
 lemma thms_excluded_middle_rule:
-  "\<lbrakk>cons(p,H) |- q;  cons(p=>Fls,H) |- q;  p \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
+  "\<lbrakk>cons(p,H) |- q;  cons(p\<Rightarrow>Fls,H) |- q;  p \<in> propn\<rbrakk> \<Longrightarrow> H |- q"
   \<comment> \<open>Hard to prove directly because it requires cuts\<close>
   apply (rule thms_excluded_middle [THEN thms_MP, THEN thms_MP])
      apply (blast intro!: propn_SIs intro: propn_Is)+
@@ -255,16 +255,16 @@
 \<close>
 
 lemma hyps_Diff:
-    "p \<in> propn \<Longrightarrow> hyps(p, t-{v}) \<subseteq> cons(#v=>Fls, hyps(p,t)-{#v})"
+    "p \<in> propn \<Longrightarrow> hyps(p, t-{v}) \<subseteq> cons(#v\<Rightarrow>Fls, hyps(p,t)-{#v})"
   by (induct set: propn) auto
 
 text \<open>
-  For the case \<^prop>\<open>hyps(p,t)-cons(#v => Fls,Y) |- p\<close> we also have
-  \<^prop>\<open>hyps(p,t)-{#v=>Fls} \<subseteq> hyps(p, cons(v,t))\<close>.
+  For the case \<^prop>\<open>hyps(p,t)-cons(#v \<Rightarrow> Fls,Y) |- p\<close> we also have
+  \<^prop>\<open>hyps(p,t)-{#v\<Rightarrow>Fls} \<subseteq> hyps(p, cons(v,t))\<close>.
 \<close>
 
 lemma hyps_cons:
-    "p \<in> propn \<Longrightarrow> hyps(p, cons(v,t)) \<subseteq> cons(#v, hyps(p,t)-{#v=>Fls})"
+    "p \<in> propn \<Longrightarrow> hyps(p, cons(v,t)) \<subseteq> cons(#v, hyps(p,t)-{#v\<Rightarrow>Fls})"
   by (induct set: propn) auto
 
 text \<open>Two lemmas for use with \<open>weaken_left\<close>\<close>
@@ -277,11 +277,11 @@
 
 text \<open>
   The set \<^term>\<open>hyps(p,t)\<close> is finite, and elements have the form
-  \<^term>\<open>#v\<close> or \<^term>\<open>#v=>Fls\<close>; could probably prove the stronger
+  \<^term>\<open>#v\<close> or \<^term>\<open>#v\<Rightarrow>Fls\<close>; could probably prove the stronger
   \<^prop>\<open>hyps(p,t) \<in> Fin(hyps(p,0) \<union> hyps(p,nat))\<close>.
 \<close>
 
-lemma hyps_finite: "p \<in> propn \<Longrightarrow> hyps(p,t) \<in> Fin(\<Union>v \<in> nat. {#v, #v=>Fls})"
+lemma hyps_finite: "p \<in> propn \<Longrightarrow> hyps(p,t) \<in> Fin(\<Union>v \<in> nat. {#v, #v\<Rightarrow>Fls})"
   by (induct set: propn) auto
 
 lemmas Diff_weaken_left = Diff_mono [OF _ subset_refl, THEN weaken_left]
@@ -304,7 +304,7 @@
     apply (blast intro: cons_Diff_same [THEN weaken_left])
    apply (blast intro: cons_Diff_subset2 [THEN weaken_left]
      hyps_Diff [THEN Diff_weaken_left])
-  txt \<open>Case \<^prop>\<open>hyps(p,t)-cons(#v => Fls,Y) |- p\<close>\<close>
+  txt \<open>Case \<^prop>\<open>hyps(p,t)-cons(#v \<Rightarrow> Fls,Y) |- p\<close>\<close>
   apply (rule thms_excluded_middle_rule)
     apply (erule_tac [3] propn.intros)
    apply (blast intro: cons_Diff_subset2 [THEN weaken_left]
@@ -321,7 +321,7 @@
   apply (blast intro: completeness_0_lemma)
   done
 
-lemma logcon_Imp: "\<lbrakk>cons(p,H) |= q\<rbrakk> \<Longrightarrow> H |= p=>q"
+lemma logcon_Imp: "\<lbrakk>cons(p,H) |= q\<rbrakk> \<Longrightarrow> H |= p\<Rightarrow>q"
   \<comment> \<open>A semantic analogue of the Deduction Theorem\<close>
   by (simp add: logcon_def)
 
--- a/src/ZF/Induct/Rmap.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Rmap.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,14 +8,14 @@
 theory Rmap imports ZF begin
 
 consts
-  rmap :: "i=>i"
+  rmap :: "i\<Rightarrow>i"
 
 inductive
   domains "rmap(r)" \<subseteq> "list(domain(r)) \<times> list(range(r))"
   intros
-    NilI:  "<Nil,Nil> \<in> rmap(r)"
+    NilI:  "\<langle>Nil,Nil\<rangle> \<in> rmap(r)"
 
-    ConsI: "\<lbrakk><x,y>: r;  <xs,ys> \<in> rmap(r)\<rbrakk>
+    ConsI: "\<lbrakk>\<langle>x,y\<rangle>: r;  \<langle>xs,ys\<rangle> \<in> rmap(r)\<rbrakk>
             \<Longrightarrow> <Cons(x,xs), Cons(y,ys)> \<in> rmap(r)"
 
   type_intros domainI rangeI list.intros
@@ -28,7 +28,7 @@
   done
 
 inductive_cases
-      Nil_rmap_case [elim!]: "<Nil,zs> \<in> rmap(r)"
+      Nil_rmap_case [elim!]: "\<langle>Nil,zs\<rangle> \<in> rmap(r)"
   and Cons_rmap_case [elim!]: "<Cons(x,xs),zs> \<in> rmap(r)"
 
 declare rmap.intros [intro]
--- a/src/ZF/Induct/Term.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Term.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,7 +12,7 @@
 \<close>
 
 consts
-  "term" :: "i => i"
+  "term" :: "i \<Rightarrow> i"
 
 datatype "term(A)" = Apply ("a \<in> A", "l \<in> list(term(A))")
   monos list_mono
@@ -21,28 +21,28 @@
 declare Apply [TC]
 
 definition
-  term_rec :: "[i, [i, i, i] => i] => i"  where
+  term_rec :: "[i, [i, i, i] \<Rightarrow> i] \<Rightarrow> i"  where
   "term_rec(t,d) \<equiv>
     Vrec(t, \<lambda>t g. term_case(\<lambda>x zs. d(x, zs, map(\<lambda>z. g`z, zs)), t))"
 
 definition
-  term_map :: "[i => i, i] => i"  where
+  term_map :: "[i \<Rightarrow> i, i] \<Rightarrow> i"  where
   "term_map(f,t) \<equiv> term_rec(t, \<lambda>x zs rs. Apply(f(x), rs))"
 
 definition
-  term_size :: "i => i"  where
+  term_size :: "i \<Rightarrow> i"  where
   "term_size(t) \<equiv> term_rec(t, \<lambda>x zs rs. succ(list_add(rs)))"
 
 definition
-  reflect :: "i => i"  where
+  reflect :: "i \<Rightarrow> i"  where
   "reflect(t) \<equiv> term_rec(t, \<lambda>x zs rs. Apply(x, rev(rs)))"
 
 definition
-  preorder :: "i => i"  where
+  preorder :: "i \<Rightarrow> i"  where
   "preorder(t) \<equiv> term_rec(t, \<lambda>x zs rs. Cons(x, flat(rs)))"
 
 definition
-  postorder :: "i => i"  where
+  postorder :: "i \<Rightarrow> i"  where
   "postorder(t) \<equiv> term_rec(t, \<lambda>x zs rs. flat(rs) @ [x])"
 
 lemma term_unfold: "term(A) = A * list(term(A))"
--- a/src/ZF/Induct/Tree_Forest.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Induct/Tree_Forest.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,9 +10,9 @@
 subsection \<open>Datatype definition\<close>
 
 consts
-  tree :: "i => i"
-  forest :: "i => i"
-  tree_forest :: "i => i"
+  tree :: "i \<Rightarrow> i"
+  forest :: "i \<Rightarrow> i"
+  tree_forest :: "i \<Rightarrow> i"
 
 datatype "tree(A)" = Tcons ("a \<in> A", "f \<in> forest(A)")
   and "forest(A)" = Fnil | Fcons ("t \<in> tree(A)", "f \<in> forest(A)")
@@ -118,12 +118,12 @@
 subsection \<open>Operations\<close>
 
 consts
-  map :: "[i => i, i] => i"
-  size :: "i => i"
-  preorder :: "i => i"
-  list_of_TF :: "i => i"
-  of_list :: "i => i"
-  reflect :: "i => i"
+  map :: "[i \<Rightarrow> i, i] \<Rightarrow> i"
+  size :: "i \<Rightarrow> i"
+  preorder :: "i \<Rightarrow> i"
+  list_of_TF :: "i \<Rightarrow> i"
+  of_list :: "i \<Rightarrow> i"
+  reflect :: "i \<Rightarrow> i"
 
 primrec
   "list_of_TF (Tcons(x,f)) = [Tcons(x,f)]"
--- a/src/ZF/Int.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Int.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,84 +10,84 @@
 definition
   intrel :: i  where
     "intrel \<equiv> {p \<in> (nat*nat)*(nat*nat).
-                \<exists>x1 y1 x2 y2. p=<<x1,y1>,<x2,y2>> \<and> x1#+y2 = x2#+y1}"
+                \<exists>x1 y1 x2 y2. p=<\<langle>x1,y1\<rangle>,\<langle>x2,y2\<rangle>> \<and> x1#+y2 = x2#+y1}"
 
 definition
   int :: i  where
     "int \<equiv> (nat*nat)//intrel"
 
 definition
-  int_of :: "i=>i" \<comment> \<open>coercion from nat to int\<close>    (\<open>$# _\<close> [80] 80)  where
+  int_of :: "i\<Rightarrow>i" \<comment> \<open>coercion from nat to int\<close>    (\<open>$# _\<close> [80] 80)  where
     "$# m \<equiv> intrel `` {<natify(m), 0>}"
 
 definition
-  intify :: "i=>i" \<comment> \<open>coercion from ANYTHING to int\<close>  where
+  intify :: "i\<Rightarrow>i" \<comment> \<open>coercion from ANYTHING to int\<close>  where
     "intify(m) \<equiv> if m \<in> int then m else $#0"
 
 definition
-  raw_zminus :: "i=>i"  where
-    "raw_zminus(z) \<equiv> \<Union><x,y>\<in>z. intrel``{<y,x>}"
+  raw_zminus :: "i\<Rightarrow>i"  where
+    "raw_zminus(z) \<equiv> \<Union>\<langle>x,y\<rangle>\<in>z. intrel``{\<langle>y,x\<rangle>}"
 
 definition
-  zminus :: "i=>i"                                 (\<open>$- _\<close> [80] 80)  where
+  zminus :: "i\<Rightarrow>i"                                 (\<open>$- _\<close> [80] 80)  where
     "$- z \<equiv> raw_zminus (intify(z))"
 
 definition
-  znegative   ::      "i=>o"  where
-    "znegative(z) \<equiv> \<exists>x y. x<y \<and> y\<in>nat \<and> <x,y>\<in>z"
+  znegative   ::      "i\<Rightarrow>o"  where
+    "znegative(z) \<equiv> \<exists>x y. x<y \<and> y\<in>nat \<and> \<langle>x,y\<rangle>\<in>z"
 
 definition
-  iszero      ::      "i=>o"  where
+  iszero      ::      "i\<Rightarrow>o"  where
     "iszero(z) \<equiv> z = $# 0"
 
 definition
-  raw_nat_of  :: "i=>i"  where
-  "raw_nat_of(z) \<equiv> natify (\<Union><x,y>\<in>z. x#-y)"
+  raw_nat_of  :: "i\<Rightarrow>i"  where
+  "raw_nat_of(z) \<equiv> natify (\<Union>\<langle>x,y\<rangle>\<in>z. x#-y)"
 
 definition
-  nat_of  :: "i=>i"  where
+  nat_of  :: "i\<Rightarrow>i"  where
   "nat_of(z) \<equiv> raw_nat_of (intify(z))"
 
 definition
-  zmagnitude  ::      "i=>i"  where
+  zmagnitude  ::      "i\<Rightarrow>i"  where
   \<comment> \<open>could be replaced by an absolute value function from int to int?\<close>
     "zmagnitude(z) \<equiv>
      THE m. m\<in>nat \<and> ((\<not> znegative(z) \<and> z = $# m) |
                        (znegative(z) \<and> $- z = $# m))"
 
 definition
-  raw_zmult   ::      "[i,i]=>i"  where
-    (*Cannot use UN<x1,y2> here or in zadd because of the form of congruent2.
+  raw_zmult   ::      "[i,i]\<Rightarrow>i"  where
+    (*Cannot use UN\<langle>x1,y2\<rangle> here or in zadd because of the form of congruent2.
       Perhaps a "curried" or even polymorphic congruent predicate would be
       better.*)
      "raw_zmult(z1,z2) \<equiv>
-       \<Union>p1\<in>z1. \<Union>p2\<in>z2.  split(%x1 y1. split(%x2 y2.
+       \<Union>p1\<in>z1. \<Union>p2\<in>z2.  split(\<lambda>x1 y1. split(\<lambda>x2 y2.
                    intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1)"
 
 definition
-  zmult       ::      "[i,i]=>i"      (infixl \<open>$*\<close> 70)  where
+  zmult       ::      "[i,i]\<Rightarrow>i"      (infixl \<open>$*\<close> 70)  where
      "z1 $* z2 \<equiv> raw_zmult (intify(z1),intify(z2))"
 
 definition
-  raw_zadd    ::      "[i,i]=>i"  where
+  raw_zadd    ::      "[i,i]\<Rightarrow>i"  where
      "raw_zadd (z1, z2) \<equiv>
-       \<Union>z1\<in>z1. \<Union>z2\<in>z2. let <x1,y1>=z1; <x2,y2>=z2
+       \<Union>z1\<in>z1. \<Union>z2\<in>z2. let \<langle>x1,y1\<rangle>=z1; \<langle>x2,y2\<rangle>=z2
                            in intrel``{<x1#+x2, y1#+y2>}"
 
 definition
-  zadd        ::      "[i,i]=>i"      (infixl \<open>$+\<close> 65)  where
+  zadd        ::      "[i,i]\<Rightarrow>i"      (infixl \<open>$+\<close> 65)  where
      "z1 $+ z2 \<equiv> raw_zadd (intify(z1),intify(z2))"
 
 definition
-  zdiff        ::      "[i,i]=>i"      (infixl \<open>$-\<close> 65)  where
+  zdiff        ::      "[i,i]\<Rightarrow>i"      (infixl \<open>$-\<close> 65)  where
      "z1 $- z2 \<equiv> z1 $+ zminus(z2)"
 
 definition
-  zless        ::      "[i,i]=>o"      (infixl \<open>$<\<close> 50)  where
+  zless        ::      "[i,i]\<Rightarrow>o"      (infixl \<open>$<\<close> 50)  where
      "z1 $< z2 \<equiv> znegative(z1 $- z2)"
 
 definition
-  zle          ::      "[i,i]=>o"      (infixl \<open>$\<le>\<close> 50)  where
+  zle          ::      "[i,i]\<Rightarrow>o"      (infixl \<open>$\<le>\<close> 50)  where
      "z1 $\<le> z2 \<equiv> z1 $< z2 | intify(z1)=intify(z2)"
 
 
@@ -98,18 +98,18 @@
 (** Natural deduction for intrel **)
 
 lemma intrel_iff [simp]:
-    "<<x1,y1>,<x2,y2>>: intrel \<longleftrightarrow>
+    "<\<langle>x1,y1\<rangle>,\<langle>x2,y2\<rangle>>: intrel \<longleftrightarrow>
      x1\<in>nat \<and> y1\<in>nat \<and> x2\<in>nat \<and> y2\<in>nat \<and> x1#+y2 = x2#+y1"
 by (simp add: intrel_def)
 
 lemma intrelI [intro!]:
     "\<lbrakk>x1#+y2 = x2#+y1; x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk>
-     \<Longrightarrow> <<x1,y1>,<x2,y2>>: intrel"
+     \<Longrightarrow> <\<langle>x1,y1\<rangle>,\<langle>x2,y2\<rangle>>: intrel"
 by (simp add: intrel_def)
 
 lemma intrelE [elim!]:
   "\<lbrakk>p \<in> intrel;
-      \<And>x1 y1 x2 y2. \<lbrakk>p = <<x1,y1>,<x2,y2>>;  x1#+y2 = x2#+y1;
+      \<And>x1 y1 x2 y2. \<lbrakk>p = <\<langle>x1,y1\<rangle>,\<langle>x2,y2\<rangle>>;  x1#+y2 = x2#+y1;
                         x1\<in>nat; y1\<in>nat; x2\<in>nat; y2\<in>nat\<rbrakk> \<Longrightarrow> Q\<rbrakk>
    \<Longrightarrow> Q"
 by (simp add: intrel_def, blast)
@@ -126,7 +126,7 @@
 apply (fast elim!: sym int_trans_lemma)
 done
 
-lemma image_intrel_int: "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> intrel `` {<m,n>} \<in> int"
+lemma image_intrel_int: "\<lbrakk>m\<in>nat; n\<in>nat\<rbrakk> \<Longrightarrow> intrel `` {\<langle>m,n\<rangle>} \<in> int"
 by (simp add: int_def)
 
 declare equiv_intrel [THEN eq_equiv_class_iff, simp]
@@ -208,7 +208,7 @@
 
 subsection\<open>\<^term>\<open>zminus\<close>: unary negation on \<^term>\<open>int\<close>\<close>
 
-lemma zminus_congruent: "(%<x,y>. intrel``{<y,x>}) respects intrel"
+lemma zminus_congruent: "(\<lambda>\<langle>x,y\<rangle>. intrel``{\<langle>y,x\<rangle>}) respects intrel"
 by (auto simp add: congruent_def add_ac)
 
 lemma raw_zminus_type: "z \<in> int \<Longrightarrow> raw_zminus(z) \<in> int"
@@ -235,13 +235,13 @@
 by auto
 
 lemma raw_zminus:
-    "\<lbrakk>x\<in>nat;  y\<in>nat\<rbrakk> \<Longrightarrow> raw_zminus(intrel``{<x,y>}) = intrel `` {<y,x>}"
+    "\<lbrakk>x\<in>nat;  y\<in>nat\<rbrakk> \<Longrightarrow> raw_zminus(intrel``{\<langle>x,y\<rangle>}) = intrel `` {\<langle>y,x\<rangle>}"
 apply (simp add: raw_zminus_def UN_equiv_class [OF equiv_intrel zminus_congruent])
 done
 
 lemma zminus:
     "\<lbrakk>x\<in>nat;  y\<in>nat\<rbrakk>
-     \<Longrightarrow> $- (intrel``{<x,y>}) = intrel `` {<y,x>}"
+     \<Longrightarrow> $- (intrel``{\<langle>x,y\<rangle>}) = intrel `` {\<langle>y,x\<rangle>}"
 by (simp add: zminus_def raw_zminus image_intrel_int)
 
 lemma raw_zminus_zminus: "z \<in> int \<Longrightarrow> raw_zminus (raw_zminus(z)) = z"
@@ -259,7 +259,7 @@
 
 subsection\<open>\<^term>\<open>znegative\<close>: the test for negative integers\<close>
 
-lemma znegative: "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> znegative(intrel``{<x,y>}) \<longleftrightarrow> x<y"
+lemma znegative: "\<lbrakk>x\<in>nat; y\<in>nat\<rbrakk> \<Longrightarrow> znegative(intrel``{\<langle>x,y\<rangle>}) \<longleftrightarrow> x<y"
 apply (cases "x<y")
 apply (auto simp add: znegative_def not_lt_iff_le)
 apply (subgoal_tac "y #+ x2 < x #+ y2", force)
@@ -286,7 +286,7 @@
 by (auto simp add: congruent_def split: nat_diff_split)
 
 lemma raw_nat_of:
-    "\<lbrakk>x\<in>nat;  y\<in>nat\<rbrakk> \<Longrightarrow> raw_nat_of(intrel``{<x,y>}) = x#-y"
+    "\<lbrakk>x\<in>nat;  y\<in>nat\<rbrakk> \<Longrightarrow> raw_nat_of(intrel``{\<langle>x,y\<rangle>}) = x#-y"
 by (simp add: raw_nat_of_def UN_equiv_class [OF equiv_intrel nat_of_congruent])
 
 lemma raw_nat_of_int_of: "raw_nat_of($# n) = natify(n)"
@@ -375,7 +375,7 @@
 
 text\<open>Congruence Property for Addition\<close>
 lemma zadd_congruent2:
-    "(%z1 z2. let <x1,y1>=z1; <x2,y2>=z2
+    "(\<lambda>z1 z2. let \<langle>x1,y1\<rangle>=z1; \<langle>x2,y2\<rangle>=z2
                             in intrel``{<x1#+x2, y1#+y2>})
      respects2 intrel"
 apply (simp add: congruent2_def)
@@ -400,7 +400,7 @@
 
 lemma raw_zadd:
   "\<lbrakk>x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat\<rbrakk>
-   \<Longrightarrow> raw_zadd (intrel``{<x1,y1>}, intrel``{<x2,y2>}) =
+   \<Longrightarrow> raw_zadd (intrel``{\<langle>x1,y1\<rangle>}, intrel``{\<langle>x2,y2\<rangle>}) =
        intrel `` {<x1#+x2, y1#+y2>}"
 apply (simp add: raw_zadd_def
              UN_equiv_class2 [OF equiv_intrel equiv_intrel zadd_congruent2])
@@ -409,7 +409,7 @@
 
 lemma zadd:
   "\<lbrakk>x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat\<rbrakk>
-   \<Longrightarrow> (intrel``{<x1,y1>}) $+ (intrel``{<x2,y2>}) =
+   \<Longrightarrow> (intrel``{\<langle>x1,y1\<rangle>}) $+ (intrel``{\<langle>x2,y2\<rangle>}) =
        intrel `` {<x1#+x2, y1#+y2>}"
 by (simp add: zadd_def raw_zadd image_intrel_int)
 
@@ -489,14 +489,14 @@
 
 text\<open>Congruence property for multiplication\<close>
 lemma zmult_congruent2:
-    "(%p1 p2. split(%x1 y1. split(%x2 y2.
+    "(\<lambda>p1 p2. split(\<lambda>x1 y1. split(\<lambda>x2 y2.
                     intrel``{<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}, p2), p1))
      respects2 intrel"
 apply (rule equiv_intrel [THEN congruent2_commuteI], auto)
 (*Proof that zmult is congruent in one argument*)
 apply (rename_tac x y)
-apply (frule_tac t = "%u. x#*u" in sym [THEN subst_context])
-apply (drule_tac t = "%u. y#*u" in subst_context)
+apply (frule_tac t = "\<lambda>u. x#*u" in sym [THEN subst_context])
+apply (drule_tac t = "\<lambda>u. y#*u" in subst_context)
 apply (erule add_left_cancel)+
 apply (simp_all add: add_mult_distrib_left)
 done
@@ -513,14 +513,14 @@
 
 lemma raw_zmult:
      "\<lbrakk>x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat\<rbrakk>
-      \<Longrightarrow> raw_zmult(intrel``{<x1,y1>}, intrel``{<x2,y2>}) =
+      \<Longrightarrow> raw_zmult(intrel``{\<langle>x1,y1\<rangle>}, intrel``{\<langle>x2,y2\<rangle>}) =
           intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
 by (simp add: raw_zmult_def
            UN_equiv_class2 [OF equiv_intrel equiv_intrel zmult_congruent2])
 
 lemma zmult:
      "\<lbrakk>x1\<in>nat; y1\<in>nat;  x2\<in>nat; y2\<in>nat\<rbrakk>
-      \<Longrightarrow> (intrel``{<x1,y1>}) $* (intrel``{<x2,y2>}) =
+      \<Longrightarrow> (intrel``{\<langle>x1,y1\<rangle>}) $* (intrel``{\<langle>x2,y2\<rangle>}) =
           intrel `` {<x1#*x2 #+ y1#*y2, x1#*y2 #+ y1#*x2>}"
 by (simp add: zmult_def raw_zmult image_intrel_int)
 
@@ -822,7 +822,7 @@
 lemma zadd_left_cancel:
      "\<lbrakk>w \<in> int; w': int\<rbrakk> \<Longrightarrow> (z $+ w' = z $+ w) \<longleftrightarrow> (w' = w)"
 apply safe
-apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
+apply (drule_tac t = "\<lambda>x. x $+ ($-z) " in subst_context)
 apply (simp add: zadd_ac)
 done
 
@@ -835,7 +835,7 @@
 lemma zadd_right_cancel:
      "\<lbrakk>w \<in> int; w': int\<rbrakk> \<Longrightarrow> (w' $+ z = w $+ z) \<longleftrightarrow> (w' = w)"
 apply safe
-apply (drule_tac t = "%x. x $+ ($-z) " in subst_context)
+apply (drule_tac t = "\<lambda>x. x $+ ($-z) " in subst_context)
 apply (simp add: zadd_ac)
 done
 
--- a/src/ZF/IntDiv.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/IntDiv.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -34,65 +34,65 @@
 begin
 
 definition
-  quorem :: "[i,i] => o"  where
-    "quorem \<equiv> %<a,b> <q,r>.
+  quorem :: "[i,i] \<Rightarrow> o"  where
+    "quorem \<equiv> \<lambda>\<langle>a,b\<rangle> \<langle>q,r\<rangle>.
                       a = b$*q $+ r \<and>
                       (#0$<b \<and> #0$\<le>r \<and> r$<b | \<not>(#0$<b) \<and> b$<r \<and> r $\<le> #0)"
 
 definition
-  adjust :: "[i,i] => i"  where
-    "adjust(b) \<equiv> %<q,r>. if #0 $\<le> r$-b then <#2$*q $+ #1,r$-b>
+  adjust :: "[i,i] \<Rightarrow> i"  where
+    "adjust(b) \<equiv> \<lambda>\<langle>q,r\<rangle>. if #0 $\<le> r$-b then <#2$*q $+ #1,r$-b>
                           else <#2$*q,r>"
 
 
 (** the division algorithm **)
 
 definition
-  posDivAlg :: "i => i"  where
+  posDivAlg :: "i \<Rightarrow> i"  where
 (*for the case a>=0, b>0*)
-(*recdef posDivAlg "inv_image less_than (%(a,b). nat_of(a $- b $+ #1))"*)
+(*recdef posDivAlg "inv_image less_than (\<lambda>\<langle>a,b\<rangle>. nat_of(a $- b $+ #1))"*)
     "posDivAlg(ab) \<equiv>
-       wfrec(measure(int*int, %<a,b>. nat_of (a $- b $+ #1)),
+       wfrec(measure(int*int, \<lambda>\<langle>a,b\<rangle>. nat_of (a $- b $+ #1)),
              ab,
-             %<a,b> f. if (a$<b | b$\<le>#0) then <#0,a>
+             \<lambda>\<langle>a,b\<rangle> f. if (a$<b | b$\<le>#0) then <#0,a>
                        else adjust(b, f ` <a,#2$*b>))"
 
 
-(*for the case a<0, b>0*)
+(*for the case a\<langle>0, b\<rangle>0*)
 definition
-  negDivAlg :: "i => i"  where
-(*recdef negDivAlg "inv_image less_than (%(a,b). nat_of(- a $- b))"*)
+  negDivAlg :: "i \<Rightarrow> i"  where
+(*recdef negDivAlg "inv_image less_than (\<lambda>\<langle>a,b\<rangle>. nat_of(- a $- b))"*)
     "negDivAlg(ab) \<equiv>
-       wfrec(measure(int*int, %<a,b>. nat_of ($- a $- b)),
+       wfrec(measure(int*int, \<lambda>\<langle>a,b\<rangle>. nat_of ($- a $- b)),
              ab,
-             %<a,b> f. if (#0 $\<le> a$+b | b$\<le>#0) then <#-1,a$+b>
+             \<lambda>\<langle>a,b\<rangle> f. if (#0 $\<le> a$+b | b$\<le>#0) then <#-1,a$+b>
                        else adjust(b, f ` <a,#2$*b>))"
 
 (*for the general case @{term"b\<noteq>0"}*)
 
 definition
-  negateSnd :: "i => i"  where
-    "negateSnd \<equiv> %<q,r>. <q, $-r>"
+  negateSnd :: "i \<Rightarrow> i"  where
+    "negateSnd \<equiv> \<lambda>\<langle>q,r\<rangle>. <q, $-r>"
 
   (*The full division algorithm considers all possible signs for a, b
     including the special case a=0, b<0, because negDivAlg requires a<0*)
 definition
-  divAlg :: "i => i"  where
+  divAlg :: "i \<Rightarrow> i"  where
     "divAlg \<equiv>
-       %<a,b>. if #0 $\<le> a then
-                  if #0 $\<le> b then posDivAlg (<a,b>)
+       \<lambda>\<langle>a,b\<rangle>. if #0 $\<le> a then
+                  if #0 $\<le> b then posDivAlg (\<langle>a,b\<rangle>)
                   else if a=#0 then <#0,#0>
                        else negateSnd (negDivAlg (<$-a,$-b>))
                else
-                  if #0$<b then negDivAlg (<a,b>)
+                  if #0$<b then negDivAlg (\<langle>a,b\<rangle>)
                   else         negateSnd (posDivAlg (<$-a,$-b>))"
 
 definition
-  zdiv  :: "[i,i]=>i"                    (infixl \<open>zdiv\<close> 70)  where
+  zdiv  :: "[i,i]\<Rightarrow>i"                    (infixl \<open>zdiv\<close> 70)  where
     "a zdiv b \<equiv> fst (divAlg (<intify(a), intify(b)>))"
 
 definition
-  zmod  :: "[i,i]=>i"                    (infixl \<open>zmod\<close> 70)  where
+  zmod  :: "[i,i]\<Rightarrow>i"                    (infixl \<open>zmod\<close> 70)  where
     "a zmod b \<equiv> snd (divAlg (<intify(a), intify(b)>))"
 
 
@@ -380,7 +380,7 @@
 
 
 lemma unique_quotient:
-     "\<lbrakk>quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \<in> int; b \<noteq> #0;
+     "\<lbrakk>quorem (\<langle>a,b\<rangle>, \<langle>q,r\<rangle>);  quorem (\<langle>a,b\<rangle>, <q',r'>);  b \<in> int; b \<noteq> #0;
          q \<in> int; q' \<in> int\<rbrakk> \<Longrightarrow> q = q'"
 apply (simp add: split_ifs quorem_def neq_iff_zless)
 apply safe
@@ -391,7 +391,7 @@
 done
 
 lemma unique_remainder:
-     "\<lbrakk>quorem (<a,b>, <q,r>);  quorem (<a,b>, <q',r'>);  b \<in> int; b \<noteq> #0;
+     "\<lbrakk>quorem (\<langle>a,b\<rangle>, \<langle>q,r\<rangle>);  quorem (\<langle>a,b\<rangle>, <q',r'>);  b \<in> int; b \<noteq> #0;
          q \<in> int; q' \<in> int;
          r \<in> int; r' \<in> int\<rbrakk> \<Longrightarrow> r = r'"
 apply (subgoal_tac "q = q'")
@@ -404,7 +404,7 @@
            the Division Algorithm for \<open>a\<ge>0\<close> and \<open>b>0\<close>\<close>
 
 lemma adjust_eq [simp]:
-     "adjust(b, <q,r>) = (let diff = r$-b in
+     "adjust(b, \<langle>q,r\<rangle>) = (let diff = r$-b in
                           if #0 $\<le> diff then <#2$*q $+ #1,diff>
                                          else <#2$*q,r>)"
 by (simp add: Let_def adjust_def)
@@ -421,7 +421,7 @@
 
 lemma posDivAlg_eqn:
      "\<lbrakk>#0 $< b; a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow>
-      posDivAlg(<a,b>) =
+      posDivAlg(\<langle>a,b\<rangle>) =
        (if a$<b then <#0,a> else adjust(b, posDivAlg (<a, #2$*b>)))"
 apply (rule posDivAlg_unfold [THEN trans])
 apply (simp add: vimage_iff not_zless_iff_zle [THEN iff_sym])
@@ -431,10 +431,10 @@
 lemma posDivAlg_induct_lemma [rule_format]:
   assumes prem:
         "\<And>a b. \<lbrakk>a \<in> int; b \<in> int;
-                   \<not> (a $< b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>)\<rbrakk> \<Longrightarrow> P(<a,b>)"
-  shows "<u,v> \<in> int*int \<Longrightarrow> P(<u,v>)"
-using wf_measure [where A = "int*int" and f = "%<a,b>.nat_of (a $- b $+ #1)"]
-proof (induct "<u,v>" arbitrary: u v rule: wf_induct)
+                   \<not> (a $< b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>)\<rbrakk> \<Longrightarrow> P(\<langle>a,b\<rangle>)"
+  shows "\<langle>u,v\<rangle> \<in> int*int \<Longrightarrow> P(\<langle>u,v\<rangle>)"
+using wf_measure [where A = "int*int" and f = "\<lambda>\<langle>a,b\<rangle>.nat_of (a $- b $+ #1)"]
+proof (induct "\<langle>u,v\<rangle>" arbitrary: u v rule: wf_induct)
   case (step x)
   hence uv: "u \<in> int" "v \<in> int" by auto
   thus ?case
@@ -452,7 +452,7 @@
       and ih: "\<And>a b. \<lbrakk>a \<in> int; b \<in> int;
                      \<not> (a $< b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b)\<rbrakk> \<Longrightarrow> P(a,b)"
   shows "P(u,v)"
-apply (subgoal_tac "(%<x,y>. P (x,y)) (<u,v>)")
+apply (subgoal_tac "(\<lambda>\<langle>x,y\<rangle>. P (x,y)) (\<langle>u,v\<rangle>)")
 apply simp
 apply (rule posDivAlg_induct_lemma)
 apply (simp (no_asm_use))
@@ -526,7 +526,7 @@
 
 (*Typechecking for posDivAlg*)
 lemma posDivAlg_type [rule_format]:
-     "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> posDivAlg(<a,b>) \<in> int * int"
+     "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> posDivAlg(\<langle>a,b\<rangle>) \<in> int * int"
 apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
 apply assumption+
 apply (case_tac "#0 $< ba")
@@ -542,7 +542,7 @@
 (*Correctness of posDivAlg: it computes quotients correctly*)
 lemma posDivAlg_correct [rule_format]:
      "\<lbrakk>a \<in> int; b \<in> int\<rbrakk>
-      \<Longrightarrow> #0 $\<le> a \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, posDivAlg(<a,b>))"
+      \<Longrightarrow> #0 $\<le> a \<longrightarrow> #0 $< b \<longrightarrow> quorem (\<langle>a,b\<rangle>, posDivAlg(\<langle>a,b\<rangle>))"
 apply (rule_tac u = "a" and v = "b" in posDivAlg_induct)
 apply auto
    apply (simp_all add: quorem_def)
@@ -576,7 +576,7 @@
 
 lemma negDivAlg_eqn:
      "\<lbrakk>#0 $< b; a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow>
-      negDivAlg(<a,b>) =
+      negDivAlg(\<langle>a,b\<rangle>) =
        (if #0 $\<le> a$+b then <#-1,a$+b>
                        else adjust(b, negDivAlg (<a, #2$*b>)))"
 apply (rule negDivAlg_unfold [THEN trans])
@@ -588,10 +588,10 @@
   assumes prem:
         "\<And>a b. \<lbrakk>a \<in> int; b \<in> int;
                    \<not> (#0 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(<a, #2 $* b>)\<rbrakk>
-                \<Longrightarrow> P(<a,b>)"
-  shows "<u,v> \<in> int*int \<Longrightarrow> P(<u,v>)"
-using wf_measure [where A = "int*int" and f = "%<a,b>.nat_of ($- a $- b)"]
-proof (induct "<u,v>" arbitrary: u v rule: wf_induct)
+                \<Longrightarrow> P(\<langle>a,b\<rangle>)"
+  shows "\<langle>u,v\<rangle> \<in> int*int \<Longrightarrow> P(\<langle>u,v\<rangle>)"
+using wf_measure [where A = "int*int" and f = "\<lambda>\<langle>a,b\<rangle>.nat_of ($- a $- b)"]
+proof (induct "\<langle>u,v\<rangle>" arbitrary: u v rule: wf_induct)
   case (step x)
   hence uv: "u \<in> int" "v \<in> int" by auto
   thus ?case
@@ -609,7 +609,7 @@
                          \<not> (#0 $\<le> a $+ b | b $\<le> #0) \<longrightarrow> P(a, #2 $* b)\<rbrakk>
                       \<Longrightarrow> P(a,b)"
   shows "P(u,v)"
-apply (subgoal_tac " (%<x,y>. P (x,y)) (<u,v>)")
+apply (subgoal_tac " (\<lambda>\<langle>x,y\<rangle>. P (x,y)) (\<langle>u,v\<rangle>)")
 apply simp
 apply (rule negDivAlg_induct_lemma)
 apply (simp (no_asm_use))
@@ -620,7 +620,7 @@
 
 (*Typechecking for negDivAlg*)
 lemma negDivAlg_type:
-     "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> negDivAlg(<a,b>) \<in> int * int"
+     "\<lbrakk>a \<in> int; b \<in> int\<rbrakk> \<Longrightarrow> negDivAlg(\<langle>a,b\<rangle>) \<in> int * int"
 apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
 apply assumption+
 apply (case_tac "#0 $< ba")
@@ -638,7 +638,7 @@
   It doesn't work if a=0 because the 0/b=0 rather than -1*)
 lemma negDivAlg_correct [rule_format]:
      "\<lbrakk>a \<in> int; b \<in> int\<rbrakk>
-      \<Longrightarrow> a $< #0 \<longrightarrow> #0 $< b \<longrightarrow> quorem (<a,b>, negDivAlg(<a,b>))"
+      \<Longrightarrow> a $< #0 \<longrightarrow> #0 $< b \<longrightarrow> quorem (\<langle>a,b\<rangle>, negDivAlg(\<langle>a,b\<rangle>))"
 apply (rule_tac u = "a" and v = "b" in negDivAlg_induct)
   apply auto
    apply (simp_all add: quorem_def)
@@ -687,7 +687,7 @@
 apply (simp add: linear_arith_lemma integ_of_type vimage_iff)
 done
 
-lemma negateSnd_eq [simp]: "negateSnd (<q,r>) = <q, $-r>"
+lemma negateSnd_eq [simp]: "negateSnd (\<langle>q,r\<rangle>) = <q, $-r>"
 apply (unfold negateSnd_def)
 apply auto
 done
@@ -699,7 +699,7 @@
 
 lemma quorem_neg:
      "\<lbrakk>quorem (<$-a,$-b>, qr);  a \<in> int;  b \<in> int;  qr \<in> int * int\<rbrakk>
-      \<Longrightarrow> quorem (<a,b>, negateSnd(qr))"
+      \<Longrightarrow> quorem (\<langle>a,b\<rangle>, negateSnd(qr))"
 apply clarify
 apply (auto elim: zless_asym simp add: quorem_def zless_zminus)
 txt\<open>linear arithmetic from here on\<close>
@@ -711,7 +711,7 @@
 done
 
 lemma divAlg_correct:
-     "\<lbrakk>b \<noteq> #0;  a \<in> int;  b \<in> int\<rbrakk> \<Longrightarrow> quorem (<a,b>, divAlg(<a,b>))"
+     "\<lbrakk>b \<noteq> #0;  a \<in> int;  b \<in> int\<rbrakk> \<Longrightarrow> quorem (\<langle>a,b\<rangle>, divAlg(\<langle>a,b\<rangle>))"
 apply (auto simp add: quorem_0 divAlg_def)
 apply (safe intro!: quorem_neg posDivAlg_correct negDivAlg_correct
                     posDivAlg_type negDivAlg_type)
@@ -720,7 +720,7 @@
 apply (auto simp add: zle_def)
 done
 
-lemma divAlg_type: "\<lbrakk>a \<in> int;  b \<in> int\<rbrakk> \<Longrightarrow> divAlg(<a,b>) \<in> int * int"
+lemma divAlg_type: "\<lbrakk>a \<in> int;  b \<in> int\<rbrakk> \<Longrightarrow> divAlg(\<langle>a,b\<rangle>) \<in> int * int"
 apply (auto simp add: divAlg_def)
 apply (auto simp add: posDivAlg_type negDivAlg_type negateSnd_type)
 done
@@ -802,20 +802,20 @@
 
 lemma quorem_div_mod:
      "\<lbrakk>b \<noteq> #0;  a \<in> int;  b \<in> int\<rbrakk>
-      \<Longrightarrow> quorem (<a,b>, <a zdiv b, a zmod b>)"
+      \<Longrightarrow> quorem (\<langle>a,b\<rangle>, <a zdiv b, a zmod b>)"
 apply (cut_tac a = "a" and b = "b" in zmod_zdiv_equality)
 apply (auto simp add: quorem_def neq_iff_zless pos_mod_sign pos_mod_bound
                       neg_mod_sign neg_mod_bound)
 done
 
-(*Surely quorem(<a,b>,<q,r>) implies @{term"a \<in> int"}, but it doesn't matter*)
+(*Surely quorem(\<langle>a,b\<rangle>,\<langle>q,r\<rangle>) implies @{term"a \<in> int"}, but it doesn't matter*)
 lemma quorem_div:
-     "\<lbrakk>quorem(<a,b>,<q,r>);  b \<noteq> #0;  a \<in> int;  b \<in> int;  q \<in> int\<rbrakk>
+     "\<lbrakk>quorem(\<langle>a,b\<rangle>,\<langle>q,r\<rangle>);  b \<noteq> #0;  a \<in> int;  b \<in> int;  q \<in> int\<rbrakk>
       \<Longrightarrow> a zdiv b = q"
 by (blast intro: quorem_div_mod [THEN unique_quotient])
 
 lemma quorem_mod:
-     "\<lbrakk>quorem(<a,b>,<q,r>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int\<rbrakk>
+     "\<lbrakk>quorem(\<langle>a,b\<rangle>,\<langle>q,r\<rangle>); b \<noteq> #0; a \<in> int; b \<in> int; q \<in> int; r \<in> int\<rbrakk>
       \<Longrightarrow> a zmod b = r"
 by (blast intro: quorem_div_mod [THEN unique_remainder])
 
@@ -953,7 +953,7 @@
 apply (simp add: int_0_less_mult_iff)
 apply (blast dest: zless_trans)
 (*linear arithmetic...*)
-apply (drule_tac t = "%x. x $- r" in subst_context)
+apply (drule_tac t = "\<lambda>x. x $- r" in subst_context)
 apply (drule sym)
 apply (simp add: zcompare_rls)
 done
@@ -963,12 +963,12 @@
  apply (simp add: int_0_le_mult_iff zcompare_rls)
  apply (blast dest: zle_zless_trans)
 apply (simp add: zdiff_zmult_distrib2)
-apply (drule_tac t = "%x. x $- a $* q" in subst_context)
+apply (drule_tac t = "\<lambda>x. x $- a $* q" in subst_context)
 apply (simp add: zcompare_rls)
 done
 
 lemma self_quotient:
-     "\<lbrakk>quorem(<a,a>,<q,r>);  a \<in> int;  q \<in> int;  a \<noteq> #0\<rbrakk> \<Longrightarrow> q = #1"
+     "\<lbrakk>quorem(\<langle>a,a\<rangle>,\<langle>q,r\<rangle>);  a \<in> int;  q \<in> int;  a \<noteq> #0\<rbrakk> \<Longrightarrow> q = #1"
 apply (simp add: split_ifs quorem_def neq_iff_zless)
 apply (rule zle_anti_sym)
 apply safe
@@ -984,7 +984,7 @@
 done
 
 lemma self_remainder:
-     "\<lbrakk>quorem(<a,a>,<q,r>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0\<rbrakk> \<Longrightarrow> r = #0"
+     "\<lbrakk>quorem(\<langle>a,a\<rangle>,\<langle>q,r\<rangle>); a \<in> int; q \<in> int; r \<in> int; a \<noteq> #0\<rbrakk> \<Longrightarrow> r = #0"
 apply (frule self_quotient)
 apply (auto simp add: quorem_def)
 done
@@ -1291,7 +1291,7 @@
 (** proving (a*b) zdiv c = a $* (b zdiv c) $+ a * (b zmod c) **)
 
 lemma zmult1_lemma:
-     "\<lbrakk>quorem(<b,c>, <q,r>);  c \<in> int;  c \<noteq> #0\<rbrakk>
+     "\<lbrakk>quorem(\<langle>b,c\<rangle>, \<langle>q,r\<rangle>);  c \<in> int;  c \<noteq> #0\<rbrakk>
       \<Longrightarrow> quorem (<a$*b, c>, <a$*q $+ (a$*r) zdiv c, (a$*r) zmod c>)"
 apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
                       pos_mod_sign pos_mod_bound neg_mod_sign neg_mod_bound)
@@ -1354,7 +1354,7 @@
             a zdiv c $+ b zdiv c $+ ((a zmod c $+ b zmod c) zdiv c) **)
 
 lemma zadd1_lemma:
-     "\<lbrakk>quorem(<a,c>, <aq,ar>);  quorem(<b,c>, <bq,br>);
+     "\<lbrakk>quorem(\<langle>a,c\<rangle>, \<langle>aq,ar\<rangle>);  quorem(\<langle>b,c\<rangle>, \<langle>bq,br\<rangle>);
          c \<in> int;  c \<noteq> #0\<rbrakk>
       \<Longrightarrow> quorem (<a$+b, c>, <aq $+ bq $+ (ar$+br) zdiv c, (ar$+br) zmod c>)"
 apply (auto simp add: split_ifs quorem_def neq_iff_zless zadd_zmult_distrib2
@@ -1511,7 +1511,7 @@
 done
 
 lemma zdiv_zmult2_lemma:
-     "\<lbrakk>quorem (<a,b>, <q,r>);  a \<in> int;  b \<in> int;  b \<noteq> #0;  #0 $< c\<rbrakk>
+     "\<lbrakk>quorem (\<langle>a,b\<rangle>, \<langle>q,r\<rangle>);  a \<in> int;  b \<in> int;  b \<noteq> #0;  #0 $< c\<rbrakk>
       \<Longrightarrow> quorem (<a,b$*c>, <q zdiv c, b$*(q zmod c) $+ r>)"
 apply (auto simp add: zmult_ac zmod_zdiv_equality [symmetric] quorem_def
                neq_iff_zless int_0_less_mult_iff
--- a/src/ZF/List.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/List.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,7 +8,7 @@
 theory List imports Datatype ArithSimp begin
 
 consts
-  list       :: "i=>i"
+  list       :: "i\<Rightarrow>i"
 
 datatype
   "list(A)" = Nil | Cons ("a \<in> A", "l \<in> list(A)")
@@ -16,7 +16,7 @@
 
 syntax
  "_Nil" :: i  (\<open>[]\<close>)
- "_List" :: "is => i"  (\<open>[(_)]\<close>)
+ "_List" :: "is \<Rightarrow> i"  (\<open>[(_)]\<close>)
 
 translations
   "[x, xs]"     == "CONST Cons(x, [xs])"
@@ -25,9 +25,9 @@
 
 
 consts
-  length :: "i=>i"
-  hd     :: "i=>i"
-  tl     :: "i=>i"
+  length :: "i\<Rightarrow>i"
+  hd     :: "i\<Rightarrow>i"
+  tl     :: "i\<Rightarrow>i"
 
 primrec
   "length([]) = 0"
@@ -43,9 +43,9 @@
 
 
 consts
-  map         :: "[i=>i, i] => i"
-  set_of_list :: "i=>i"
-  app         :: "[i,i]=>i"                        (infixr \<open>@\<close> 60)
+  map         :: "[i\<Rightarrow>i, i] \<Rightarrow> i"
+  set_of_list :: "i\<Rightarrow>i"
+  app         :: "[i,i]\<Rightarrow>i"                        (infixr \<open>@\<close> 60)
 
 (*map is a binding operator -- it applies to meta-level functions, not
 object-level functions.  This simplifies the final form of term_rec_conv,
@@ -64,9 +64,9 @@
 
 
 consts
-  rev :: "i=>i"
-  flat       :: "i=>i"
-  list_add   :: "i=>i"
+  rev :: "i\<Rightarrow>i"
+  flat       :: "i\<Rightarrow>i"
+  list_add   :: "i\<Rightarrow>i"
 
 primrec
   "rev([]) = []"
@@ -81,7 +81,7 @@
   "list_add(Cons(a,l)) = a #+ list_add(l)"
 
 consts
-  drop       :: "[i,i]=>i"
+  drop       :: "[i,i]\<Rightarrow>i"
 
 primrec
   drop_0:    "drop(0,l) = l"
@@ -92,25 +92,25 @@
 
 definition
 (* Function `take' returns the first n elements of a list *)
-  take     :: "[i,i]=>i"  where
+  take     :: "[i,i]\<Rightarrow>i"  where
   "take(n, as) \<equiv> list_rec(\<lambda>n\<in>nat. [],
-                %a l r. \<lambda>n\<in>nat. nat_case([], %m. Cons(a, r`m), n), as)`n"
+                \<lambda>a l r. \<lambda>n\<in>nat. nat_case([], \<lambda>m. Cons(a, r`m), n), as)`n"
 
 definition
-  nth :: "[i, i]=>i"  where
+  nth :: "[i, i]\<Rightarrow>i"  where
   \<comment> \<open>returns the (n+1)th element of a list, or 0 if the
    list is too short.\<close>
   "nth(n, as) \<equiv> list_rec(\<lambda>n\<in>nat. 0,
-                          %a l r. \<lambda>n\<in>nat. nat_case(a, %m. r`m, n), as) ` n"
+                          \<lambda>a l r. \<lambda>n\<in>nat. nat_case(a, \<lambda>m. r`m, n), as) ` n"
 
 definition
-  list_update :: "[i, i, i]=>i"  where
+  list_update :: "[i, i, i]\<Rightarrow>i"  where
   "list_update(xs, i, v) \<equiv> list_rec(\<lambda>n\<in>nat. Nil,
-      %u us vs. \<lambda>n\<in>nat. nat_case(Cons(v, us), %m. Cons(u, vs`m), n), xs)`i"
+      \<lambda>u us vs. \<lambda>n\<in>nat. nat_case(Cons(v, us), \<lambda>m. Cons(u, vs`m), n), xs)`i"
 
 consts
-  filter :: "[i=>o, i] => i"
-  upt :: "[i, i] =>i"
+  filter :: "[i\<Rightarrow>o, i] \<Rightarrow> i"
+  upt :: "[i, i] \<Rightarrow>i"
 
 primrec
   "filter(P, Nil) = Nil"
@@ -122,11 +122,11 @@
   "upt(i, succ(j)) = (if i \<le> j then upt(i, j)@[j] else Nil)"
 
 definition
-  min :: "[i,i] =>i"  where
+  min :: "[i,i] \<Rightarrow>i"  where
     "min(x, y) \<equiv> (if x \<le> y then x else y)"
 
 definition
-  max :: "[i, i] =>i"  where
+  max :: "[i, i] \<Rightarrow>i"  where
     "max(x, y) \<equiv> (if x \<le> y then y else x)"
 
 (*** Aspects of the datatype definition ***)
@@ -286,12 +286,12 @@
 
 (*** theorems about map ***)
 
-lemma map_ident [simp]: "l \<in> list(A) \<Longrightarrow> map(%u. u, l) = l"
+lemma map_ident [simp]: "l \<in> list(A) \<Longrightarrow> map(\<lambda>u. u, l) = l"
 apply (induct_tac "l")
 apply (simp_all (no_asm_simp))
 done
 
-lemma map_compose: "l \<in> list(A) \<Longrightarrow> map(h, map(j,l)) = map(%u. h(j(u)), l)"
+lemma map_compose: "l \<in> list(A) \<Longrightarrow> map(h, map(j,l)) = map(\<lambda>u. h(j(u)), l)"
 apply (induct_tac "l")
 apply (simp_all (no_asm_simp))
 done
@@ -309,7 +309,7 @@
 lemma list_rec_map:
      "l \<in> list(A) \<Longrightarrow>
       list_rec(c, d, map(h,l)) =
-      list_rec(c, %x xs r. d(h(x), map(h,xs), r), l)"
+      list_rec(c, \<lambda>x xs r. d(h(x), map(h,xs), r), l)"
 apply (induct_tac "l")
 apply (simp_all (no_asm_simp))
 done
@@ -510,10 +510,10 @@
 lemma filter_is_subset: "xs:list(A) \<Longrightarrow> set_of_list(filter(P,xs)) \<subseteq> set_of_list(xs)"
 by (induct_tac "xs", auto)
 
-lemma filter_False [simp]: "xs:list(A) \<Longrightarrow> filter(%p. False, xs) = Nil"
+lemma filter_False [simp]: "xs:list(A) \<Longrightarrow> filter(\<lambda>p. False, xs) = Nil"
 by (induct_tac "xs", auto)
 
-lemma filter_True [simp]: "xs:list(A) \<Longrightarrow> filter(%p. True, xs) = xs"
+lemma filter_True [simp]: "xs:list(A) \<Longrightarrow> filter(\<lambda>p. True, xs) = xs"
 by (induct_tac "xs", auto)
 
 (** length **)
@@ -853,18 +853,18 @@
 text\<open>Crafty definition to eliminate a type argument\<close>
 
 consts
-  zip_aux        :: "[i,i]=>i"
+  zip_aux        :: "[i,i]\<Rightarrow>i"
 
 primrec (*explicit lambda is required because both arguments of "un" vary*)
   "zip_aux(B,[]) =
-     (\<lambda>ys \<in> list(B). list_case([], %y l. [], ys))"
+     (\<lambda>ys \<in> list(B). list_case([], \<lambda>y l. [], ys))"
 
   "zip_aux(B,Cons(x,l)) =
      (\<lambda>ys \<in> list(B).
-        list_case(Nil, %y zs. Cons(<x,y>, zip_aux(B,l)`zs), ys))"
+        list_case(Nil, \<lambda>y zs. Cons(\<langle>x,y\<rangle>, zip_aux(B,l)`zs), ys))"
 
 definition
-  zip :: "[i, i]=>i"  where
+  zip :: "[i, i]\<Rightarrow>i"  where
    "zip(xs, ys) \<equiv> zip_aux(set_of_list(ys),xs)`ys"
 
 
@@ -897,7 +897,7 @@
 
 lemma zip_Cons_Cons [simp]:
      "\<lbrakk>xs:list(A); ys:list(B); x \<in> A; y \<in> B\<rbrakk> \<Longrightarrow>
-      zip(Cons(x,xs), Cons(y, ys)) = Cons(<x,y>, zip(xs, ys))"
+      zip(Cons(x,xs), Cons(y, ys)) = Cons(\<langle>x,y\<rangle>, zip(xs, ys))"
 apply (simp add: zip_def, auto)
 apply (rule zip_aux_unique, auto)
 apply (simp add: list_on_set_of_list [of _ B])
@@ -967,7 +967,7 @@
 lemma set_of_list_zip [rule_format]:
      "\<lbrakk>xs:list(A); ys:list(B); i \<in> nat\<rbrakk>
       \<Longrightarrow> set_of_list(zip(xs, ys)) =
-          {<x, y>:A*B. \<exists>i\<in>nat. i < min(length(xs), length(ys))
+          {\<langle>x, y\<rangle>:A*B. \<exists>i\<in>nat. i < min(length(xs), length(ys))
           \<and> x = nth(i, xs) \<and> y = nth(i, ys)}"
 by (force intro!: Collect_cong simp add: lt_min_iff set_of_list_conv_nth)
 
@@ -1179,9 +1179,9 @@
 (** sublist (a generalization of nth to sets) **)
 
 definition
-  sublist :: "[i, i] => i"  where
+  sublist :: "[i, i] \<Rightarrow> i"  where
     "sublist(xs, A)\<equiv>
-     map(fst, (filter(%p. snd(p): A, zip(xs, upt(0,length(xs))))))"
+     map(fst, (filter(\<lambda>p. snd(p): A, zip(xs, upt(0,length(xs))))))"
 
 lemma sublist_0 [simp]: "xs:list(A) \<Longrightarrow>sublist(xs, 0) =Nil"
 by (unfold sublist_def, auto)
@@ -1191,8 +1191,8 @@
 
 lemma sublist_shift_lemma:
  "\<lbrakk>xs:list(B); i \<in> nat\<rbrakk> \<Longrightarrow>
-  map(fst, filter(%p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) =
-  map(fst, filter(%p. snd(p):nat \<and> snd(p) #+ i \<in> A, zip(xs,upt(0,length(xs)))))"
+  map(fst, filter(\<lambda>p. snd(p):A, zip(xs, upt(i,i #+ length(xs))))) =
+  map(fst, filter(\<lambda>p. snd(p):nat \<and> snd(p) #+ i \<in> A, zip(xs,upt(0,length(xs)))))"
 apply (erule list_append_induct)
 apply (simp (no_asm_simp))
 apply (auto simp add: add_commute length_app filter_append map_app_distrib)
@@ -1250,7 +1250,7 @@
 
 text\<open>Repetition of a List Element\<close>
 
-consts   repeat :: "[i,i]=>i"
+consts   repeat :: "[i,i]\<Rightarrow>i"
 primrec
   "repeat(a,0) = []"
 
--- a/src/ZF/Nat.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Nat.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,49 +9,49 @@
 
 definition
   nat :: i  where
-    "nat \<equiv> lfp(Inf, %X. {0} \<union> {succ(i). i \<in> X})"
+    "nat \<equiv> lfp(Inf, \<lambda>X. {0} \<union> {succ(i). i \<in> X})"
 
 definition
-  quasinat :: "i => o"  where
+  quasinat :: "i \<Rightarrow> o"  where
     "quasinat(n) \<equiv> n=0 | (\<exists>m. n = succ(m))"
 
 definition
   (*Has an unconditional succ case, which is used in "recursor" below.*)
-  nat_case :: "[i, i=>i, i]=>i"  where
+  nat_case :: "[i, i\<Rightarrow>i, i]\<Rightarrow>i"  where
     "nat_case(a,b,k) \<equiv> THE y. k=0 \<and> y=a | (\<exists>x. k=succ(x) \<and> y=b(x))"
 
 definition
-  nat_rec :: "[i, i, [i,i]=>i]=>i"  where
+  nat_rec :: "[i, i, [i,i]\<Rightarrow>i]\<Rightarrow>i"  where
     "nat_rec(k,a,b) \<equiv>
-          wfrec(Memrel(nat), k, %n f. nat_case(a, %m. b(m, f`m), n))"
+          wfrec(Memrel(nat), k, \<lambda>n f. nat_case(a, \<lambda>m. b(m, f`m), n))"
 
   (*Internalized relations on the naturals*)
 
 definition
   Le :: i  where
-    "Le \<equiv> {<x,y>:nat*nat. x \<le> y}"
+    "Le \<equiv> {\<langle>x,y\<rangle>:nat*nat. x \<le> y}"
 
 definition
   Lt :: i  where
-    "Lt \<equiv> {<x, y>:nat*nat. x < y}"
+    "Lt \<equiv> {\<langle>x, y\<rangle>:nat*nat. x < y}"
 
 definition
   Ge :: i  where
-    "Ge \<equiv> {<x,y>:nat*nat. y \<le> x}"
+    "Ge \<equiv> {\<langle>x,y\<rangle>:nat*nat. y \<le> x}"
 
 definition
   Gt :: i  where
-    "Gt \<equiv> {<x,y>:nat*nat. y < x}"
+    "Gt \<equiv> {\<langle>x,y\<rangle>:nat*nat. y < x}"
 
 definition
-  greater_than :: "i=>i"  where
+  greater_than :: "i\<Rightarrow>i"  where
     "greater_than(n) \<equiv> {i \<in> nat. n < i}"
 
 text\<open>No need for a less-than operator: a natural number is its list of
 predecessors!\<close>
 
 
-lemma nat_bnd_mono: "bnd_mono(Inf, %X. {0} \<union> {succ(i). i \<in> X})"
+lemma nat_bnd_mono: "bnd_mono(Inf, \<lambda>X. {0} \<union> {succ(i). i \<in> X})"
 apply (rule bnd_monoI)
 apply (cut_tac infinity, blast, blast)
 done
@@ -291,7 +291,7 @@
 apply (blast intro: lt_trans)
 done
 
-lemma Le_iff [iff]: "<x,y> \<in> Le \<longleftrightarrow> x \<le> y \<and> x \<in> nat \<and> y \<in> nat"
+lemma Le_iff [iff]: "\<langle>x,y\<rangle> \<in> Le \<longleftrightarrow> x \<le> y \<and> x \<in> nat \<and> y \<in> nat"
 by (force simp add: Le_def)
 
 end
--- a/src/ZF/OrdQuant.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/OrdQuant.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,22 +10,22 @@
 
 definition
   (* Ordinal Quantifiers *)
-  oall :: "[i, i => o] => o"  where
+  oall :: "[i, i \<Rightarrow> o] \<Rightarrow> o"  where
     "oall(A, P) \<equiv> \<forall>x. x<A \<longrightarrow> P(x)"
 
 definition
-  oex :: "[i, i => o] => o"  where
+  oex :: "[i, i \<Rightarrow> o] \<Rightarrow> o"  where
     "oex(A, P)  \<equiv> \<exists>x. x<A \<and> P(x)"
 
 definition
   (* Ordinal Union *)
-  OUnion :: "[i, i => i] => i"  where
+  OUnion :: "[i, i \<Rightarrow> i] \<Rightarrow> i"  where
     "OUnion(i,B) \<equiv> {z: \<Union>x\<in>i. B(x). Ord(i)}"
 
 syntax
-  "_oall"     :: "[idt, i, o] => o"        (\<open>(3\<forall>_<_./ _)\<close> 10)
-  "_oex"      :: "[idt, i, o] => o"        (\<open>(3\<exists>_<_./ _)\<close> 10)
-  "_OUNION"   :: "[idt, i, i] => i"        (\<open>(3\<Union>_<_./ _)\<close> 10)
+  "_oall"     :: "[idt, i, o] \<Rightarrow> o"        (\<open>(3\<forall>_<_./ _)\<close> 10)
+  "_oex"      :: "[idt, i, o] \<Rightarrow> o"        (\<open>(3\<exists>_<_./ _)\<close> 10)
+  "_OUNION"   :: "[idt, i, i] \<Rightarrow> i"        (\<open>(3\<Union>_<_./ _)\<close> 10)
 translations
   "\<forall>x<a. P" \<rightleftharpoons> "CONST oall(a, \<lambda>x. P)"
   "\<exists>x<a. P" \<rightleftharpoons> "CONST oex(a, \<lambda>x. P)"
@@ -127,7 +127,7 @@
 (*Congruence rule for rewriting*)
 lemma oall_cong [cong]:
     "\<lbrakk>a=a';  \<And>x. x<a' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk>
-     \<Longrightarrow> oall(a, %x. P(x)) <-> oall(a', %x. P'(x))"
+     \<Longrightarrow> oall(a, \<lambda>x. P(x)) <-> oall(a', \<lambda>x. P'(x))"
 by (simp add: oall_def)
 
 
@@ -151,7 +151,7 @@
 
 lemma oex_cong [cong]:
     "\<lbrakk>a=a';  \<And>x. x<a' \<Longrightarrow> P(x) <-> P'(x)\<rbrakk>
-     \<Longrightarrow> oex(a, %x. P(x)) <-> oex(a', %x. P'(x))"
+     \<Longrightarrow> oex(a, \<lambda>x. P(x)) <-> oex(a', \<lambda>x. P'(x))"
 apply (simp add: oex_def cong add: conj_cong)
 done
 
@@ -184,16 +184,16 @@
 subsection \<open>Quantification over a class\<close>
 
 definition
-  "rall"     :: "[i=>o, i=>o] => o"  where
+  "rall"     :: "[i\<Rightarrow>o, i\<Rightarrow>o] \<Rightarrow> o"  where
     "rall(M, P) \<equiv> \<forall>x. M(x) \<longrightarrow> P(x)"
 
 definition
-  "rex"      :: "[i=>o, i=>o] => o"  where
+  "rex"      :: "[i\<Rightarrow>o, i\<Rightarrow>o] \<Rightarrow> o"  where
     "rex(M, P) \<equiv> \<exists>x. M(x) \<and> P(x)"
 
 syntax
-  "_rall"     :: "[pttrn, i=>o, o] => o"        (\<open>(3\<forall>_[_]./ _)\<close> 10)
-  "_rex"      :: "[pttrn, i=>o, o] => o"        (\<open>(3\<exists>_[_]./ _)\<close> 10)
+  "_rall"     :: "[pttrn, i\<Rightarrow>o, o] \<Rightarrow> o"        (\<open>(3\<forall>_[_]./ _)\<close> 10)
+  "_rex"      :: "[pttrn, i\<Rightarrow>o, o] \<Rightarrow> o"        (\<open>(3\<exists>_[_]./ _)\<close> 10)
 translations
   "\<forall>x[M]. P" \<rightleftharpoons> "CONST rall(M, \<lambda>x. P)"
   "\<exists>x[M]. P" \<rightleftharpoons> "CONST rex(M, \<lambda>x. P)"
@@ -249,10 +249,10 @@
     "(\<And>x. M(x) \<Longrightarrow> P(x) <-> P'(x)) \<Longrightarrow> (\<exists>x[M]. P(x)) <-> (\<exists>x[M]. P'(x))"
 by (simp add: rex_def cong: conj_cong)
 
-lemma rall_is_ball [simp]: "(\<forall>x[%z. z\<in>A]. P(x)) <-> (\<forall>x\<in>A. P(x))"
+lemma rall_is_ball [simp]: "(\<forall>x[\<lambda>z. z\<in>A]. P(x)) <-> (\<forall>x\<in>A. P(x))"
 by blast
 
-lemma rex_is_bex [simp]: "(\<exists>x[%z. z\<in>A]. P(x)) <-> (\<exists>x\<in>A. P(x))"
+lemma rex_is_bex [simp]: "(\<exists>x[\<lambda>z. z\<in>A]. P(x)) <-> (\<exists>x\<in>A. P(x))"
 by blast
 
 lemma atomize_rall: "(\<And>x. M(x) \<Longrightarrow> P(x)) \<equiv> Trueprop (\<forall>x[M]. P(x))"
@@ -323,8 +323,8 @@
 subsubsection\<open>Sets as Classes\<close>
 
 definition
-  setclass :: "[i,i] => o"       (\<open>##_\<close> [40] 40)  where
-   "setclass(A) \<equiv> %x. x \<in> A"
+  setclass :: "[i,i] \<Rightarrow> o"       (\<open>##_\<close> [40] 40)  where
+   "setclass(A) \<equiv> \<lambda>x. x \<in> A"
 
 lemma setclass_iff [simp]: "setclass(A,x) <-> x \<in> A"
 by (simp add: setclass_def)
--- a/src/ZF/Order.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Order.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -16,15 +16,15 @@
   counterparts.\<close>
 
 definition
-  part_ord :: "[i,i]=>o"                (*Strict partial ordering*)  where
+  part_ord :: "[i,i]\<Rightarrow>o"                (*Strict partial ordering*)  where
    "part_ord(A,r) \<equiv> irrefl(A,r) \<and> trans[A](r)"
 
 definition
-  linear   :: "[i,i]=>o"                (*Strict total ordering*)  where
-   "linear(A,r) \<equiv> (\<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r | x=y | <y,x>:r)"
+  linear   :: "[i,i]\<Rightarrow>o"                (*Strict total ordering*)  where
+   "linear(A,r) \<equiv> (\<forall>x\<in>A. \<forall>y\<in>A. \<langle>x,y\<rangle>:r | x=y | \<langle>y,x\<rangle>:r)"
 
 definition
-  tot_ord  :: "[i,i]=>o"                (*Strict total ordering*)  where
+  tot_ord  :: "[i,i]\<Rightarrow>o"                (*Strict total ordering*)  where
    "tot_ord(A,r) \<equiv> part_ord(A,r) \<and> linear(A,r)"
 
 definition
@@ -40,31 +40,31 @@
   "Partial_order(r) \<equiv> partial_order_on(field(r), r)"
 
 definition
-  well_ord :: "[i,i]=>o"                (*Well-ordering*)  where
+  well_ord :: "[i,i]\<Rightarrow>o"                (*Well-ordering*)  where
    "well_ord(A,r) \<equiv> tot_ord(A,r) \<and> wf[A](r)"
 
 definition
-  mono_map :: "[i,i,i,i]=>i"            (*Order-preserving maps*)  where
+  mono_map :: "[i,i,i,i]\<Rightarrow>i"            (*Order-preserving maps*)  where
    "mono_map(A,r,B,s) \<equiv>
-              {f \<in> A->B. \<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r \<longrightarrow> <f`x,f`y>:s}"
+              {f \<in> A->B. \<forall>x\<in>A. \<forall>y\<in>A. \<langle>x,y\<rangle>:r \<longrightarrow> <f`x,f`y>:s}"
 
 definition
-  ord_iso  :: "[i,i,i,i]=>i"  (\<open>(\<langle>_, _\<rangle> \<cong>/ \<langle>_, _\<rangle>)\<close> 51)  (*Order isomorphisms*)  where
+  ord_iso  :: "[i,i,i,i]\<Rightarrow>i"  (\<open>(\<langle>_, _\<rangle> \<cong>/ \<langle>_, _\<rangle>)\<close> 51)  (*Order isomorphisms*)  where
    "\<langle>A,r\<rangle> \<cong> \<langle>B,s\<rangle> \<equiv>
-              {f \<in> bij(A,B). \<forall>x\<in>A. \<forall>y\<in>A. <x,y>:r \<longleftrightarrow> <f`x,f`y>:s}"
+              {f \<in> bij(A,B). \<forall>x\<in>A. \<forall>y\<in>A. \<langle>x,y\<rangle>:r \<longleftrightarrow> <f`x,f`y>:s}"
 
 definition
-  pred     :: "[i,i,i]=>i"              (*Set of predecessors*)  where
-   "pred(A,x,r) \<equiv> {y \<in> A. <y,x>:r}"
+  pred     :: "[i,i,i]\<Rightarrow>i"              (*Set of predecessors*)  where
+   "pred(A,x,r) \<equiv> {y \<in> A. \<langle>y,x\<rangle>:r}"
 
 definition
-  ord_iso_map :: "[i,i,i,i]=>i"         (*Construction for linearity theorem*)  where
+  ord_iso_map :: "[i,i,i,i]\<Rightarrow>i"         (*Construction for linearity theorem*)  where
    "ord_iso_map(A,r,B,s) \<equiv>
-     \<Union>x\<in>A. \<Union>y\<in>B. \<Union>f \<in> ord_iso(pred(A,x,r), r, pred(B,y,s), s). {<x,y>}"
+     \<Union>x\<in>A. \<Union>y\<in>B. \<Union>f \<in> ord_iso(pred(A,x,r), r, pred(B,y,s), s). {\<langle>x,y\<rangle>}"
 
 definition
-  first :: "[i, i, i] => o"  where
-    "first(u, X, R) \<equiv> u \<in> X \<and> (\<forall>v\<in>X. v\<noteq>u \<longrightarrow> <u,v> \<in> R)"
+  first :: "[i, i, i] \<Rightarrow> o"  where
+    "first(u, X, R) \<equiv> u \<in> X \<and> (\<forall>v\<in>X. v\<noteq>u \<longrightarrow> \<langle>u,v\<rangle> \<in> R)"
 
 subsection\<open>Immediate Consequences of the Definitions\<close>
 
@@ -74,7 +74,7 @@
 
 lemma linearE:
     "\<lbrakk>linear(A,r);  x \<in> A;  y \<in> A;
-        <x,y>:r \<Longrightarrow> P;  x=y \<Longrightarrow> P;  <y,x>:r \<Longrightarrow> P\<rbrakk>
+        \<langle>x,y\<rangle>:r \<Longrightarrow> P;  x=y \<Longrightarrow> P;  \<langle>y,x\<rangle>:r \<Longrightarrow> P\<rbrakk>
      \<Longrightarrow> P"
 by (simp add: linear_def, blast)
 
@@ -102,12 +102,12 @@
 
 (** Derived rules for pred(A,x,r) **)
 
-lemma pred_iff: "y \<in> pred(A,x,r) \<longleftrightarrow> <y,x>:r \<and> y \<in> A"
+lemma pred_iff: "y \<in> pred(A,x,r) \<longleftrightarrow> \<langle>y,x\<rangle>:r \<and> y \<in> A"
 by (unfold pred_def, blast)
 
 lemmas predI = conjI [THEN pred_iff [THEN iffD2]]
 
-lemma predE: "\<lbrakk>y \<in> pred(A,x,r);  \<lbrakk>y \<in> A; <y,x>:r\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+lemma predE: "\<lbrakk>y \<in> pred(A,x,r);  \<lbrakk>y \<in> A; \<langle>y,x\<rangle>:r\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (simp add: pred_def)
 
 lemma pred_subset_under: "pred(A,x,r) \<subseteq> r -`` {x}"
@@ -121,7 +121,7 @@
 by (simp add: pred_def, blast)
 
 lemma trans_pred_pred_eq:
-    "\<lbrakk>trans[A](r);  <y,x>:r;  x \<in> A;  y \<in> A\<rbrakk>
+    "\<lbrakk>trans[A](r);  \<langle>y,x\<rangle>:r;  x \<in> A;  y \<in> A\<rbrakk>
      \<Longrightarrow> pred(pred(A,x,r), y, r) = pred(A,y,r)"
 by (unfold trans_on_def pred_def, blast)
 
@@ -251,7 +251,7 @@
 
 lemma ord_isoI:
     "\<lbrakk>f \<in> bij(A, B);
-        \<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> <x, y> \<in> r \<longleftrightarrow> <f`x, f`y> \<in> s\<rbrakk>
+        \<And>x y. \<lbrakk>x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> \<langle>x, y\<rangle> \<in> r \<longleftrightarrow> <f`x, f`y> \<in> s\<rbrakk>
      \<Longrightarrow> f \<in> ord_iso(A,r,B,s)"
 by (simp add: ord_iso_def)
 
@@ -267,11 +267,11 @@
 
 (*Needed?  But ord_iso_converse is!*)
 lemma ord_iso_apply:
-    "\<lbrakk>f \<in> ord_iso(A,r,B,s);  <x,y>: r;  x \<in> A;  y \<in> A\<rbrakk> \<Longrightarrow> <f`x, f`y> \<in> s"
+    "\<lbrakk>f \<in> ord_iso(A,r,B,s);  \<langle>x,y\<rangle>: r;  x \<in> A;  y \<in> A\<rbrakk> \<Longrightarrow> <f`x, f`y> \<in> s"
 by (simp add: ord_iso_def)
 
 lemma ord_iso_converse:
-    "\<lbrakk>f \<in> ord_iso(A,r,B,s);  <x,y>: s;  x \<in> B;  y \<in> B\<rbrakk>
+    "\<lbrakk>f \<in> ord_iso(A,r,B,s);  \<langle>x,y\<rangle>: s;  x \<in> B;  y \<in> B\<rbrakk>
      \<Longrightarrow> <converse(f) ` x, converse(f) ` y> \<in> r"
 apply (simp add: ord_iso_def, clarify)
 apply (erule bspec [THEN bspec, THEN iffD2])
@@ -437,10 +437,10 @@
 
 (*Tricky; a lot of forward proof!*)
 lemma well_ord_iso_preserving:
-     "\<lbrakk>well_ord(A,r);  well_ord(B,s);  <a,c>: r;
+     "\<lbrakk>well_ord(A,r);  well_ord(B,s);  \<langle>a,c\<rangle>: r;
          f \<in> ord_iso(pred(A,a,r), r, pred(B,b,s), s);
          g \<in> ord_iso(pred(A,c,r), r, pred(B,d,s), s);
-         a \<in> A;  c \<in> A;  b \<in> B;  d \<in> B\<rbrakk> \<Longrightarrow> <b,d>: s"
+         a \<in> A;  c \<in> A;  b \<in> B;  d \<in> B\<rbrakk> \<Longrightarrow> \<langle>b,d\<rangle>: s"
 apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type], (erule asm_rl predI predE)+)
 apply (subgoal_tac "b = g`a")
 apply (simp (no_asm_simp))
@@ -552,7 +552,7 @@
 (*Trivial case: b=a*)
 apply clarify
 apply blast
-(*Harder case: <a, xa>: r*)
+(*Harder case: \<langle>a, xa\<rangle>: r*)
 apply (frule ord_iso_is_bij [THEN bij_is_fun, THEN apply_type],
        (erule asm_rl predI predE)+)
 apply (frule ord_iso_restrict_pred)
@@ -603,7 +603,7 @@
 apply (rule wf_on_not_refl [THEN notE])
   apply (erule well_ord_is_wf)
  apply assumption
-apply (subgoal_tac "<x,y>: ord_iso_map (A,r,B,s) ")
+apply (subgoal_tac "\<langle>x,y\<rangle>: ord_iso_map (A,r,B,s) ")
  apply (drule rangeI)
  apply (simp add: pred_def)
 apply (unfold ord_iso_map_def, blast)
@@ -661,7 +661,7 @@
 
 lemma subset_vimage_vimage_iff:
   "\<lbrakk>Preorder(r); A \<subseteq> field(r); B \<subseteq> field(r)\<rbrakk> \<Longrightarrow>
-  r -`` A \<subseteq> r -`` B \<longleftrightarrow> (\<forall>a\<in>A. \<exists>b\<in>B. <a, b> \<in> r)"
+  r -`` A \<subseteq> r -`` B \<longleftrightarrow> (\<forall>a\<in>A. \<exists>b\<in>B. \<langle>a, b\<rangle> \<in> r)"
   apply (auto simp: subset_def preorder_on_def refl_def vimage_def image_def)
    apply blast
   unfolding trans_on_def
@@ -674,7 +674,7 @@
 
 lemma subset_vimage1_vimage1_iff:
   "\<lbrakk>Preorder(r); a \<in> field(r); b \<in> field(r)\<rbrakk> \<Longrightarrow>
-  r -`` {a} \<subseteq> r -`` {b} \<longleftrightarrow> <a, b> \<in> r"
+  r -`` {a} \<subseteq> r -`` {b} \<longleftrightarrow> \<langle>a, b\<rangle> \<in> r"
   by (simp add: subset_vimage_vimage_iff)
 
 lemma Refl_antisym_eq_Image1_Image1_iff:
--- a/src/ZF/OrderArith.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/OrderArith.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,29 +9,29 @@
 
 definition
   (*disjoint sum of two relations; underlies ordinal addition*)
-  radd    :: "[i,i,i,i]=>i"  where
+  radd    :: "[i,i,i,i]\<Rightarrow>i"  where
     "radd(A,r,B,s) \<equiv>
                 {z: (A+B) * (A+B).
-                    (\<exists>x y. z = <Inl(x), Inr(y)>)   |
-                    (\<exists>x' x. z = <Inl(x'), Inl(x)> \<and> <x',x>:r)   |
-                    (\<exists>y' y. z = <Inr(y'), Inr(y)> \<and> <y',y>:s)}"
+                    (\<exists>x y. z = \<langle>Inl(x), Inr(y)\<rangle>)   |
+                    (\<exists>x' x. z = \<langle>Inl(x'), Inl(x)\<rangle> \<and> \<langle>x',x\<rangle>:r)   |
+                    (\<exists>y' y. z = \<langle>Inr(y'), Inr(y)\<rangle> \<and> \<langle>y',y\<rangle>:s)}"
 
 definition
   (*lexicographic product of two relations; underlies ordinal multiplication*)
-  rmult   :: "[i,i,i,i]=>i"  where
+  rmult   :: "[i,i,i,i]\<Rightarrow>i"  where
     "rmult(A,r,B,s) \<equiv>
                 {z: (A*B) * (A*B).
-                    \<exists>x' y' x y. z = <<x',y'>, <x,y>> \<and>
-                       (<x',x>: r | (x'=x \<and> <y',y>: s))}"
+                    \<exists>x' y' x y. z = \<langle>\<langle>x',y'\<rangle>, \<langle>x,y\<rangle>\<rangle> \<and>
+                       (\<langle>x',x\<rangle>: r | (x'=x \<and> \<langle>y',y\<rangle>: s))}"
 
 definition
   (*inverse image of a relation*)
-  rvimage :: "[i,i,i]=>i"  where
-    "rvimage(A,f,r) \<equiv> {z \<in> A*A. \<exists>x y. z = <x,y> \<and> <f`x,f`y>: r}"
+  rvimage :: "[i,i,i]\<Rightarrow>i"  where
+    "rvimage(A,f,r) \<equiv> {z \<in> A*A. \<exists>x y. z = \<langle>x,y\<rangle> \<and> \<langle>f`x,f`y\<rangle>: r}"
 
 definition
   measure :: "[i, i\<Rightarrow>i] \<Rightarrow> i"  where
-    "measure(A,f) \<equiv> {<x,y>: A*A. f(x) < f(y)}"
+    "measure(A,f) \<equiv> {\<langle>x,y\<rangle>: A*A. f(x) < f(y)}"
 
 
 subsection\<open>Addition of Relations -- Disjoint Sum\<close>
@@ -39,19 +39,19 @@
 subsubsection\<open>Rewrite rules.  Can be used to obtain introduction rules\<close>
 
 lemma radd_Inl_Inr_iff [iff]:
-    "<Inl(a), Inr(b)> \<in> radd(A,r,B,s)  \<longleftrightarrow>  a \<in> A \<and> b \<in> B"
+    "\<langle>Inl(a), Inr(b)\<rangle> \<in> radd(A,r,B,s)  \<longleftrightarrow>  a \<in> A \<and> b \<in> B"
 by (unfold radd_def, blast)
 
 lemma radd_Inl_iff [iff]:
-    "<Inl(a'), Inl(a)> \<in> radd(A,r,B,s)  \<longleftrightarrow>  a':A \<and> a \<in> A \<and> <a',a>:r"
+    "\<langle>Inl(a'), Inl(a)\<rangle> \<in> radd(A,r,B,s)  \<longleftrightarrow>  a':A \<and> a \<in> A \<and> \<langle>a',a\<rangle>:r"
 by (unfold radd_def, blast)
 
 lemma radd_Inr_iff [iff]:
-    "<Inr(b'), Inr(b)> \<in> radd(A,r,B,s) \<longleftrightarrow>  b':B \<and> b \<in> B \<and> <b',b>:s"
+    "\<langle>Inr(b'), Inr(b)\<rangle> \<in> radd(A,r,B,s) \<longleftrightarrow>  b':B \<and> b \<in> B \<and> \<langle>b',b\<rangle>:s"
 by (unfold radd_def, blast)
 
 lemma radd_Inr_Inl_iff [simp]:
-    "<Inr(b), Inl(a)> \<in> radd(A,r,B,s) \<longleftrightarrow> False"
+    "\<langle>Inr(b), Inl(a)\<rangle> \<in> radd(A,r,B,s) \<longleftrightarrow> False"
 by (unfold radd_def, blast)
 
 declare radd_Inr_Inl_iff [THEN iffD1, dest!]
@@ -59,10 +59,10 @@
 subsubsection\<open>Elimination Rule\<close>
 
 lemma raddE:
-    "\<lbrakk><p',p> \<in> radd(A,r,B,s);
+    "\<lbrakk>\<langle>p',p\<rangle> \<in> radd(A,r,B,s);
         \<And>x y. \<lbrakk>p'=Inl(x); x \<in> A; p=Inr(y); y \<in> B\<rbrakk> \<Longrightarrow> Q;
-        \<And>x' x. \<lbrakk>p'=Inl(x'); p=Inl(x); <x',x>: r; x':A; x \<in> A\<rbrakk> \<Longrightarrow> Q;
-        \<And>y' y. \<lbrakk>p'=Inr(y'); p=Inr(y); <y',y>: s; y':B; y \<in> B\<rbrakk> \<Longrightarrow> Q
+        \<And>x' x. \<lbrakk>p'=Inl(x'); p=Inl(x); \<langle>x',x\<rangle>: r; x':A; x \<in> A\<rbrakk> \<Longrightarrow> Q;
+        \<And>y' y. \<lbrakk>p'=Inr(y'); p=Inr(y); \<langle>y',y\<rangle>: s; y':B; y \<in> B\<rbrakk> \<Longrightarrow> Q
 \<rbrakk> \<Longrightarrow> Q"
 by (unfold radd_def, blast)
 
@@ -116,8 +116,8 @@
 
 lemma sum_bij:
      "\<lbrakk>f \<in> bij(A,C);  g \<in> bij(B,D)\<rbrakk>
-      \<Longrightarrow> (\<lambda>z\<in>A+B. case(%x. Inl(f`x), %y. Inr(g`y), z)) \<in> bij(A+B, C+D)"
-apply (rule_tac d = "case (%x. Inl (converse(f)`x), %y. Inr(converse(g)`y))"
+      \<Longrightarrow> (\<lambda>z\<in>A+B. case(\<lambda>x. Inl(f`x), \<lambda>y. Inr(g`y), z)) \<in> bij(A+B, C+D)"
+apply (rule_tac d = "case (\<lambda>x. Inl (converse(f)`x), \<lambda>y. Inr(converse(g)`y))"
        in lam_bijective)
 apply (typecheck add: bij_is_inj inj_is_fun)
 apply (auto simp add: left_inverse_bij right_inverse_bij)
@@ -125,7 +125,7 @@
 
 lemma sum_ord_iso_cong:
     "\<lbrakk>f \<in> ord_iso(A,r,A',r');  g \<in> ord_iso(B,s,B',s')\<rbrakk> \<Longrightarrow>
-            (\<lambda>z\<in>A+B. case(%x. Inl(f`x), %y. Inr(g`y), z))
+            (\<lambda>z\<in>A+B. case(\<lambda>x. Inl(f`x), \<lambda>y. Inr(g`y), z))
             \<in> ord_iso(A+B, radd(A,r,B,s), A'+B', radd(A',r',B',s'))"
 apply (unfold ord_iso_def)
 apply (safe intro!: sum_bij)
@@ -136,23 +136,23 @@
 (*Could we prove an ord_iso result?  Perhaps
      ord_iso(A+B, radd(A,r,B,s), A \<union> B, r \<union> s) *)
 lemma sum_disjoint_bij: "A \<inter> B = 0 \<Longrightarrow>
-            (\<lambda>z\<in>A+B. case(%x. x, %y. y, z)) \<in> bij(A+B, A \<union> B)"
-apply (rule_tac d = "%z. if z \<in> A then Inl (z) else Inr (z) " in lam_bijective)
+            (\<lambda>z\<in>A+B. case(\<lambda>x. x, \<lambda>y. y, z)) \<in> bij(A+B, A \<union> B)"
+apply (rule_tac d = "\<lambda>z. if z \<in> A then Inl (z) else Inr (z) " in lam_bijective)
 apply auto
 done
 
 subsubsection\<open>Associativity\<close>
 
 lemma sum_assoc_bij:
-     "(\<lambda>z\<in>(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z))
+     "(\<lambda>z\<in>(A+B)+C. case(case(Inl, \<lambda>y. Inr(Inl(y))), \<lambda>y. Inr(Inr(y)), z))
       \<in> bij((A+B)+C, A+(B+C))"
-apply (rule_tac d = "case (%x. Inl (Inl (x)), case (%x. Inl (Inr (x)), Inr))"
+apply (rule_tac d = "case (\<lambda>x. Inl (Inl (x)), case (\<lambda>x. Inl (Inr (x)), Inr))"
        in lam_bijective)
 apply auto
 done
 
 lemma sum_assoc_ord_iso:
-     "(\<lambda>z\<in>(A+B)+C. case(case(Inl, %y. Inr(Inl(y))), %y. Inr(Inr(y)), z))
+     "(\<lambda>z\<in>(A+B)+C. case(case(Inl, \<lambda>y. Inr(Inl(y))), \<lambda>y. Inr(Inr(y)), z))
       \<in> ord_iso((A+B)+C, radd(A+B, radd(A,r,B,s), C, t),
                 A+(B+C), radd(A, r, B+C, radd(B,s,C,t)))"
 by (rule sum_assoc_bij [THEN ord_isoI], auto)
@@ -163,16 +163,16 @@
 subsubsection\<open>Rewrite rule.  Can be used to obtain introduction rules\<close>
 
 lemma  rmult_iff [iff]:
-    "<<a',b'>, <a,b>> \<in> rmult(A,r,B,s) \<longleftrightarrow>
-            (<a',a>: r  \<and> a':A \<and> a \<in> A \<and> b': B \<and> b \<in> B) |
-            (<b',b>: s  \<and> a'=a \<and> a \<in> A \<and> b': B \<and> b \<in> B)"
+    "\<langle>\<langle>a',b'\<rangle>, \<langle>a,b\<rangle>\<rangle> \<in> rmult(A,r,B,s) \<longleftrightarrow>
+            (\<langle>a',a\<rangle>: r  \<and> a':A \<and> a \<in> A \<and> b': B \<and> b \<in> B) |
+            (\<langle>b',b\<rangle>: s  \<and> a'=a \<and> a \<in> A \<and> b': B \<and> b \<in> B)"
 
 by (unfold rmult_def, blast)
 
 lemma rmultE:
-    "\<lbrakk><<a',b'>, <a,b>> \<in> rmult(A,r,B,s);
-        \<lbrakk><a',a>: r;  a':A;  a \<in> A;  b':B;  b \<in> B\<rbrakk> \<Longrightarrow> Q;
-        \<lbrakk><b',b>: s;  a \<in> A;  a'=a;  b':B;  b \<in> B\<rbrakk> \<Longrightarrow> Q
+    "\<lbrakk>\<langle>\<langle>a',b'\<rangle>, \<langle>a,b\<rangle>\<rangle> \<in> rmult(A,r,B,s);
+        \<lbrakk>\<langle>a',a\<rangle>: r;  a':A;  a \<in> A;  b':B;  b \<in> B\<rbrakk> \<Longrightarrow> Q;
+        \<lbrakk>\<langle>b',b\<rangle>: s;  a \<in> A;  a'=a;  b':B;  b \<in> B\<rbrakk> \<Longrightarrow> Q
 \<rbrakk> \<Longrightarrow> Q"
 by blast
 
@@ -195,7 +195,7 @@
 apply (rule wf_onI2)
 apply (erule SigmaE)
 apply (erule ssubst)
-apply (subgoal_tac "\<forall>b\<in>B. <x,b>: Ba", blast)
+apply (subgoal_tac "\<forall>b\<in>B. \<langle>x,b\<rangle>: Ba", blast)
 apply (erule_tac a = x in wf_on_induct, assumption)
 apply (rule ballI)
 apply (erule_tac a = b in wf_on_induct, assumption)
@@ -221,8 +221,8 @@
 
 lemma prod_bij:
      "\<lbrakk>f \<in> bij(A,C);  g \<in> bij(B,D)\<rbrakk>
-      \<Longrightarrow> (lam <x,y>:A*B. <f`x, g`y>) \<in> bij(A*B, C*D)"
-apply (rule_tac d = "%<x,y>. <converse (f) `x, converse (g) `y>"
+      \<Longrightarrow> (lam \<langle>x,y\<rangle>:A*B. \<langle>f`x, g`y\<rangle>) \<in> bij(A*B, C*D)"
+apply (rule_tac d = "\<lambda>\<langle>x,y\<rangle>. \<langle>converse (f) `x, converse (g) `y\<rangle>"
        in lam_bijective)
 apply (typecheck add: bij_is_inj inj_is_fun)
 apply (auto simp add: left_inverse_bij right_inverse_bij)
@@ -230,7 +230,7 @@
 
 lemma prod_ord_iso_cong:
     "\<lbrakk>f \<in> ord_iso(A,r,A',r');  g \<in> ord_iso(B,s,B',s')\<rbrakk>
-     \<Longrightarrow> (lam <x,y>:A*B. <f`x, g`y>)
+     \<Longrightarrow> (lam \<langle>x,y\<rangle>:A*B. \<langle>f`x, g`y\<rangle>)
          \<in> ord_iso(A*B, rmult(A,r,B,s), A'*B', rmult(A',r',B',s'))"
 apply (unfold ord_iso_def)
 apply (safe intro!: prod_bij)
@@ -238,13 +238,13 @@
 apply (blast intro: bij_is_inj [THEN inj_apply_equality])
 done
 
-lemma singleton_prod_bij: "(\<lambda>z\<in>A. <x,z>) \<in> bij(A, {x}*A)"
+lemma singleton_prod_bij: "(\<lambda>z\<in>A. \<langle>x,z\<rangle>) \<in> bij(A, {x}*A)"
 by (rule_tac d = snd in lam_bijective, auto)
 
 (*Used??*)
 lemma singleton_prod_ord_iso:
      "well_ord({x},xr) \<Longrightarrow>
-          (\<lambda>z\<in>A. <x,z>) \<in> ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))"
+          (\<lambda>z\<in>A. \<langle>x,z\<rangle>) \<in> ord_iso(A, r, {x}*A, rmult({x}, xr, A, r))"
 apply (rule singleton_prod_bij [THEN ord_isoI])
 apply (simp (no_asm_simp))
 apply (blast dest: well_ord_is_wf [THEN wf_on_not_refl])
@@ -254,7 +254,7 @@
   case_cong, id_conv, comp_lam, case_case.*)
 lemma prod_sum_singleton_bij:
      "a\<notin>C \<Longrightarrow>
-       (\<lambda>x\<in>C*B + D. case(%x. x, %y.<a,y>, x))
+       (\<lambda>x\<in>C*B + D. case(\<lambda>x. x, \<lambda>y.\<langle>a,y\<rangle>, x))
        \<in> bij(C*B + D, C*B \<union> {a}*D)"
 apply (rule subst_elem)
 apply (rule id_bij [THEN sum_bij, THEN comp_bij])
@@ -268,7 +268,7 @@
 
 lemma prod_sum_singleton_ord_iso:
  "\<lbrakk>a \<in> A;  well_ord(A,r)\<rbrakk> \<Longrightarrow>
-    (\<lambda>x\<in>pred(A,a,r)*B + pred(B,b,s). case(%x. x, %y.<a,y>, x))
+    (\<lambda>x\<in>pred(A,a,r)*B + pred(B,b,s). case(\<lambda>x. x, \<lambda>y.\<langle>a,y\<rangle>, x))
     \<in> ord_iso(pred(A,a,r)*B + pred(B,b,s),
                   radd(A*B, rmult(A,r,B,s), B, s),
               pred(A,a,r)*B \<union> {a}*pred(B,b,s), rmult(A,r,B,s))"
@@ -280,13 +280,13 @@
 subsubsection\<open>Distributive law\<close>
 
 lemma sum_prod_distrib_bij:
-     "(lam <x,z>:(A+B)*C. case(%y. Inl(<y,z>), %y. Inr(<y,z>), x))
+     "(lam \<langle>x,z\<rangle>:(A+B)*C. case(\<lambda>y. Inl(\<langle>y,z\<rangle>), \<lambda>y. Inr(\<langle>y,z\<rangle>), x))
       \<in> bij((A+B)*C, (A*C)+(B*C))"
-by (rule_tac d = "case (%<x,y>.<Inl (x),y>, %<x,y>.<Inr (x),y>) "
+by (rule_tac d = "case (\<lambda>\<langle>x,y\<rangle>.\<langle>Inl (x),y\<rangle>, \<lambda>\<langle>x,y\<rangle>.\<langle>Inr (x),y\<rangle>) "
     in lam_bijective, auto)
 
 lemma sum_prod_distrib_ord_iso:
- "(lam <x,z>:(A+B)*C. case(%y. Inl(<y,z>), %y. Inr(<y,z>), x))
+ "(lam \<langle>x,z\<rangle>:(A+B)*C. case(\<lambda>y. Inl(\<langle>y,z\<rangle>), \<lambda>y. Inr(\<langle>y,z\<rangle>), x))
   \<in> ord_iso((A+B)*C, rmult(A+B, radd(A,r,B,s), C, t),
             (A*C)+(B*C), radd(A*C, rmult(A,r,C,t), B*C, rmult(B,s,C,t)))"
 by (rule sum_prod_distrib_bij [THEN ord_isoI], auto)
@@ -294,11 +294,11 @@
 subsubsection\<open>Associativity\<close>
 
 lemma prod_assoc_bij:
-     "(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>) \<in> bij((A*B)*C, A*(B*C))"
-by (rule_tac d = "%<x, <y,z>>. <<x,y>, z>" in lam_bijective, auto)
+     "(lam \<langle>\<langle>x,y\<rangle>, z\<rangle>:(A*B)*C. \<langle>x,\<langle>y,z\<rangle>\<rangle>) \<in> bij((A*B)*C, A*(B*C))"
+by (rule_tac d = "\<lambda>\<langle>x, \<langle>y,z\<rangle>\<rangle>. \<langle>\<langle>x,y\<rangle>, z\<rangle>" in lam_bijective, auto)
 
 lemma prod_assoc_ord_iso:
- "(lam <<x,y>, z>:(A*B)*C. <x,<y,z>>)
+ "(lam \<langle>\<langle>x,y\<rangle>, z\<rangle>:(A*B)*C. \<langle>x,\<langle>y,z\<rangle>\<rangle>)
   \<in> ord_iso((A*B)*C, rmult(A*B, rmult(A,r,B,s), C, t),
             A*(B*C), rmult(A, r, B*C, rmult(B,s,C,t)))"
 by (rule prod_assoc_bij [THEN ord_isoI], auto)
@@ -307,7 +307,7 @@
 
 subsubsection\<open>Rewrite rule\<close>
 
-lemma rvimage_iff: "<a,b> \<in> rvimage(A,f,r)  \<longleftrightarrow>  <f`a,f`b>: r \<and> a \<in> A \<and> b \<in> A"
+lemma rvimage_iff: "\<langle>a,b\<rangle> \<in> rvimage(A,f,r)  \<longleftrightarrow>  \<langle>f`a,f`b\<rangle>: r \<and> a \<in> A \<and> b \<in> A"
 by (unfold rvimage_def, blast)
 
 subsubsection\<open>Type checking\<close>
@@ -371,7 +371,7 @@
 
 text\<open>But note that the combination of \<open>wf_imp_wf_on\<close> and
  \<open>wf_rvimage\<close> gives \<^prop>\<open>wf(r) \<Longrightarrow> wf[C](rvimage(A,f,r))\<close>\<close>
-lemma wf_on_rvimage: "\<lbrakk>f \<in> A->B;  wf[B](r)\<rbrakk> \<Longrightarrow> wf[A](rvimage(A,f,r))"
+lemma wf_on_rvimage: "\<lbrakk>f \<in> A\<rightarrow>B;  wf[B](r)\<rbrakk> \<Longrightarrow> wf[A](rvimage(A,f,r))"
 apply (rule wf_onI2)
 apply (subgoal_tac "\<forall>z\<in>A. f`z=f`y \<longrightarrow> z \<in> Ba")
  apply blast
@@ -408,11 +408,11 @@
 
 
 definition
-  wfrank :: "[i,i]=>i"  where
-    "wfrank(r,a) \<equiv> wfrec(r, a, %x f. \<Union>y \<in> r-``{x}. succ(f`y))"
+  wfrank :: "[i,i]\<Rightarrow>i"  where
+    "wfrank(r,a) \<equiv> wfrec(r, a, \<lambda>x f. \<Union>y \<in> r-``{x}. succ(f`y))"
 
 definition
-  wftype :: "i=>i"  where
+  wftype :: "i\<Rightarrow>i"  where
     "wftype(r) \<equiv> \<Union>y \<in> range(r). succ(wfrank(r,y))"
 
 lemma wfrank: "wf(r) \<Longrightarrow> wfrank(r,a) = (\<Union>y \<in> r-``{a}. succ(wfrank(r,y)))"
@@ -424,7 +424,7 @@
 apply (rule Ord_succ [THEN Ord_UN], blast)
 done
 
-lemma wfrank_lt: "\<lbrakk>wf(r); <a,b> \<in> r\<rbrakk> \<Longrightarrow> wfrank(r,a) < wfrank(r,b)"
+lemma wfrank_lt: "\<lbrakk>wf(r); \<langle>a,b\<rangle> \<in> r\<rbrakk> \<Longrightarrow> wfrank(r,a) < wfrank(r,b)"
 apply (rule_tac a1 = b in wfrank [THEN ssubst], assumption)
 apply (rule UN_I [THEN ltI])
 apply (simp add: Ord_wfrank vimage_iff)+
@@ -496,7 +496,7 @@
 lemma wf_measure [iff]: "wf(measure(A,f))"
 by (simp (no_asm) add: measure_eq_rvimage_Memrel wf_Memrel wf_rvimage)
 
-lemma measure_iff [iff]: "<x,y> \<in> measure(A,f) \<longleftrightarrow> x \<in> A \<and> y \<in> A \<and> f(x)<f(y)"
+lemma measure_iff [iff]: "\<langle>x,y\<rangle> \<in> measure(A,f) \<longleftrightarrow> x \<in> A \<and> y \<in> A \<and> f(x)<f(y)"
 by (simp (no_asm) add: measure_def)
 
 lemma linear_measure:
@@ -529,8 +529,8 @@
 lemma wf_on_Union:
  assumes wfA: "wf[A](r)"
      and wfB: "\<And>a. a\<in>A \<Longrightarrow> wf[B(a)](s)"
-     and ok: "\<And>a u v. \<lbrakk><u,v> \<in> s; v \<in> B(a); a \<in> A\<rbrakk>
-                       \<Longrightarrow> (\<exists>a'\<in>A. <a',a> \<in> r \<and> u \<in> B(a')) | u \<in> B(a)"
+     and ok: "\<And>a u v. \<lbrakk>\<langle>u,v\<rangle> \<in> s; v \<in> B(a); a \<in> A\<rbrakk>
+                       \<Longrightarrow> (\<exists>a'\<in>A. \<langle>a',a\<rangle> \<in> r \<and> u \<in> B(a')) | u \<in> B(a)"
  shows "wf[\<Union>a\<in>A. B(a)](s)"
 apply (rule wf_onI2)
 apply (erule UN_E)
@@ -547,18 +547,18 @@
 subsubsection\<open>Bijections involving Powersets\<close>
 
 lemma Pow_sum_bij:
-    "(\<lambda>Z \<in> Pow(A+B). <{x \<in> A. Inl(x) \<in> Z}, {y \<in> B. Inr(y) \<in> Z}>)
+    "(\<lambda>Z \<in> Pow(A+B). \<langle>{x \<in> A. Inl(x) \<in> Z}, {y \<in> B. Inr(y) \<in> Z}\<rangle>)
      \<in> bij(Pow(A+B), Pow(A)*Pow(B))"
-apply (rule_tac d = "%<X,Y>. {Inl (x). x \<in> X} \<union> {Inr (y). y \<in> Y}"
+apply (rule_tac d = "\<lambda>\<langle>X,Y\<rangle>. {Inl (x). x \<in> X} \<union> {Inr (y). y \<in> Y}"
        in lam_bijective)
 apply force+
 done
 
-text\<open>As a special case, we have \<^term>\<open>bij(Pow(A*B), A -> Pow(B))\<close>\<close>
+text\<open>As a special case, we have \<^term>\<open>bij(Pow(A*B), A \<rightarrow> Pow(B))\<close>\<close>
 lemma Pow_Sigma_bij:
     "(\<lambda>r \<in> Pow(Sigma(A,B)). \<lambda>x \<in> A. r``{x})
      \<in> bij(Pow(Sigma(A,B)), \<Prod>x \<in> A. Pow(B(x)))"
-apply (rule_tac d = "%f. \<Union>x \<in> A. \<Union>y \<in> f`x. {<x,y>}" in lam_bijective)
+apply (rule_tac d = "\<lambda>f. \<Union>x \<in> A. \<Union>y \<in> f`x. {\<langle>x,y\<rangle>}" in lam_bijective)
 apply (blast intro: lam_type)
 apply (blast dest: apply_type, simp_all)
 apply fast (*strange, but blast can't do it*)
--- a/src/ZF/OrderType.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/OrderType.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,40 +12,40 @@
 here.  But a definition by transfinite recursion would be much simpler!\<close>
 
 definition
-  ordermap  :: "[i,i]=>i"  where
-   "ordermap(A,r) \<equiv> \<lambda>x\<in>A. wfrec[A](r, x, %x f. f `` pred(A,x,r))"
+  ordermap  :: "[i,i]\<Rightarrow>i"  where
+   "ordermap(A,r) \<equiv> \<lambda>x\<in>A. wfrec[A](r, x, \<lambda>x f. f `` pred(A,x,r))"
 
 definition
-  ordertype :: "[i,i]=>i"  where
+  ordertype :: "[i,i]\<Rightarrow>i"  where
    "ordertype(A,r) \<equiv> ordermap(A,r)``A"
 
 definition
   (*alternative definition of ordinal numbers*)
-  Ord_alt   :: "i => o"  where
+  Ord_alt   :: "i \<Rightarrow> o"  where
    "Ord_alt(X) \<equiv> well_ord(X, Memrel(X)) \<and> (\<forall>u\<in>X. u=pred(X, u, Memrel(X)))"
 
 definition
   (*coercion to ordinal: if not, just 0*)
-  ordify    :: "i=>i"  where
+  ordify    :: "i\<Rightarrow>i"  where
     "ordify(x) \<equiv> if Ord(x) then x else 0"
 
 definition
   (*ordinal multiplication*)
-  omult      :: "[i,i]=>i"           (infixl \<open>**\<close> 70)  where
+  omult      :: "[i,i]\<Rightarrow>i"           (infixl \<open>**\<close> 70)  where
    "i ** j \<equiv> ordertype(j*i, rmult(j,Memrel(j),i,Memrel(i)))"
 
 definition
   (*ordinal addition*)
-  raw_oadd   :: "[i,i]=>i"  where
+  raw_oadd   :: "[i,i]\<Rightarrow>i"  where
     "raw_oadd(i,j) \<equiv> ordertype(i+j, radd(i,Memrel(i),j,Memrel(j)))"
 
 definition
-  oadd      :: "[i,i]=>i"           (infixl \<open>++\<close> 65)  where
+  oadd      :: "[i,i]\<Rightarrow>i"           (infixl \<open>++\<close> 65)  where
     "i ++ j \<equiv> raw_oadd(ordify(i),ordify(j))"
 
 definition
   (*ordinal subtraction*)
-  odiff      :: "[i,i]=>i"           (infixl \<open>--\<close> 65)  where
+  odiff      :: "[i,i]\<Rightarrow>i"           (infixl \<open>--\<close> 65)  where
     "i -- j \<equiv> ordertype(i-j, Memrel(i))"
 
 
@@ -128,14 +128,14 @@
 (*The theorem above is
 
 \<lbrakk>wf[A](r); x \<in> A\<rbrakk>
-\<Longrightarrow> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \<in> A . <y,x> \<in> r}}
+\<Longrightarrow> ordermap(A,r) ` x = {ordermap(A,r) ` y . y: {y \<in> A . \<langle>y,x\<rangle> \<in> r}}
 
 NOTE: the definition of ordermap used here delivers ordinals only if r is
 transitive.  If r is the predecessor relation on the naturals then
 ordermap(nat,predr) ` n equals {n-1} and not n.  A more complicated definition,
 like
 
-  ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y \<in> A . <y,x> \<in> r}},
+  ordermap(A,r) ` x = Union{succ(ordermap(A,r) ` y) . y: {y \<in> A . \<langle>y,x\<rangle> \<in> r}},
 
 might eliminate the need for r to be transitive.
 *)
@@ -169,7 +169,7 @@
 subsubsection\<open>ordermap preserves the orderings in both directions\<close>
 
 lemma ordermap_mono:
-     "\<lbrakk><w,x>: r;  wf[A](r);  w \<in> A; x \<in> A\<rbrakk>
+     "\<lbrakk>\<langle>w,x\<rangle>: r;  wf[A](r);  w \<in> A; x \<in> A\<rbrakk>
       \<Longrightarrow> ordermap(A,r)`w \<in> ordermap(A,r)`x"
 apply (erule_tac x1 = x in ordermap_unfold [THEN ssubst], assumption, blast)
 done
@@ -177,7 +177,7 @@
 (*linearity of r is crucial here*)
 lemma converse_ordermap_mono:
     "\<lbrakk>ordermap(A,r)`w \<in> ordermap(A,r)`x;  well_ord(A,r); w \<in> A; x \<in> A\<rbrakk>
-     \<Longrightarrow> <w,x>: r"
+     \<Longrightarrow> \<langle>w,x\<rangle>: r"
 apply (unfold well_ord_def tot_ord_def, safe)
 apply (erule_tac x=w and y=x in linearE, assumption+)
 apply (blast elim!: mem_not_refl [THEN notE])
@@ -333,7 +333,7 @@
 
 text\<open>Addition with 0\<close>
 
-lemma bij_sum_0: "(\<lambda>z\<in>A+0. case(%x. x, %y. y, z)) \<in> bij(A+0, A)"
+lemma bij_sum_0: "(\<lambda>z\<in>A+0. case(\<lambda>x. x, \<lambda>y. y, z)) \<in> bij(A+0, A)"
 apply (rule_tac d = Inl in lam_bijective, safe)
 apply (simp_all (no_asm_simp))
 done
@@ -345,7 +345,7 @@
 apply force
 done
 
-lemma bij_0_sum: "(\<lambda>z\<in>0+A. case(%x. x, %y. y, z)) \<in> bij(0+A, A)"
+lemma bij_0_sum: "(\<lambda>z\<in>0+A. case(\<lambda>x. x, \<lambda>y. y, z)) \<in> bij(0+A, A)"
 apply (rule_tac d = Inr in lam_bijective, safe)
 apply (simp_all (no_asm_simp))
 done
@@ -364,7 +364,7 @@
  "a \<in> A \<Longrightarrow> (\<lambda>x\<in>pred(A,a,r). Inl(x))
           \<in> bij(pred(A,a,r), pred(A+B, Inl(a), radd(A,r,B,s)))"
 apply (unfold pred_def)
-apply (rule_tac d = "case (%x. x, %y. y) " in lam_bijective)
+apply (rule_tac d = "case (\<lambda>x. x, \<lambda>y. y) " in lam_bijective)
 apply auto
 done
 
@@ -382,7 +382,7 @@
          id(A+pred(B,b,s))
          \<in> bij(A+pred(B,b,s), pred(A+B, Inr(b), radd(A,r,B,s)))"
 apply (unfold pred_def id_def)
-apply (rule_tac d = "%z. z" in lam_bijective, auto)
+apply (rule_tac d = "\<lambda>z. z" in lam_bijective, auto)
 done
 
 lemma ordertype_pred_Inr_eq:
@@ -686,7 +686,7 @@
 
 lemma bij_sum_Diff:
      "A<=B \<Longrightarrow> (\<lambda>y\<in>B. if(y \<in> A, Inl(y), Inr(y))) \<in> bij(B, A+(B-A))"
-apply (rule_tac d = "case (%x. x, %y. y) " in lam_bijective)
+apply (rule_tac d = "case (\<lambda>x. x, \<lambda>y. y) " in lam_bijective)
 apply (blast intro!: if_type)
 apply (fast intro!: case_type)
 apply (erule_tac [2] sumE)
@@ -757,14 +757,14 @@
 subsubsection\<open>A useful unfolding law\<close>
 
 lemma pred_Pair_eq:
- "\<lbrakk>a \<in> A;  b \<in> B\<rbrakk> \<Longrightarrow> pred(A*B, <a,b>, rmult(A,r,B,s)) =
+ "\<lbrakk>a \<in> A;  b \<in> B\<rbrakk> \<Longrightarrow> pred(A*B, \<langle>a,b\<rangle>, rmult(A,r,B,s)) =
                       pred(A,a,r)*B \<union> ({a} * pred(B,b,s))"
 apply (unfold pred_def, blast)
 done
 
 lemma ordertype_pred_Pair_eq:
      "\<lbrakk>a \<in> A;  b \<in> B;  well_ord(A,r);  well_ord(B,s)\<rbrakk> \<Longrightarrow>
-         ordertype(pred(A*B, <a,b>, rmult(A,r,B,s)), rmult(A,r,B,s)) =
+         ordertype(pred(A*B, \<langle>a,b\<rangle>, rmult(A,r,B,s)), rmult(A,r,B,s)) =
          ordertype(pred(A,a,r)*B + pred(B,b,s),
                   radd(A*B, rmult(A,r,B,s), B, s))"
 apply (simp (no_asm_simp) add: pred_Pair_eq)
@@ -849,7 +849,7 @@
 apply (unfold omult_def)
 apply (rule_tac s1="Memrel(i)"
        in ord_isoI [THEN ordertype_eq, THEN trans])
-apply (rule_tac c = snd and d = "%z.<0,z>"  in lam_bijective)
+apply (rule_tac c = snd and d = "\<lambda>z.\<langle>0,z\<rangle>"  in lam_bijective)
 apply (auto elim!: snd_type well_ord_Memrel ordertype_Memrel)
 done
 
@@ -857,7 +857,7 @@
 apply (unfold omult_def)
 apply (rule_tac s1="Memrel(i)"
        in ord_isoI [THEN ordertype_eq, THEN trans])
-apply (rule_tac c = fst and d = "%z.<z,0>" in lam_bijective)
+apply (rule_tac c = fst and d = "\<lambda>z.\<langle>z,0\<rangle>" in lam_bijective)
 apply (auto elim!: fst_type well_ord_Memrel ordertype_Memrel)
 done
 
--- a/src/ZF/Ordinal.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Ordinal.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,23 +8,23 @@
 theory Ordinal imports WF Bool equalities begin
 
 definition
-  Memrel        :: "i=>i"  where
-    "Memrel(A)   \<equiv> {z\<in>A*A . \<exists>x y. z=<x,y> \<and> x\<in>y }"
+  Memrel        :: "i\<Rightarrow>i"  where
+    "Memrel(A)   \<equiv> {z\<in>A*A . \<exists>x y. z=\<langle>x,y\<rangle> \<and> x\<in>y }"
 
 definition
-  Transset  :: "i=>o"  where
+  Transset  :: "i\<Rightarrow>o"  where
     "Transset(i) \<equiv> \<forall>x\<in>i. x<=i"
 
 definition
-  Ord  :: "i=>o"  where
+  Ord  :: "i\<Rightarrow>o"  where
     "Ord(i)      \<equiv> Transset(i) \<and> (\<forall>x\<in>i. Transset(x))"
 
 definition
-  lt        :: "[i,i] => o"  (infixl \<open><\<close> 50)   (*less-than on ordinals*)  where
+  lt        :: "[i,i] \<Rightarrow> o"  (infixl \<open><\<close> 50)   (*less-than on ordinals*)  where
     "i<j         \<equiv> i\<in>j \<and> Ord(j)"
 
 definition
-  Limit         :: "i=>o"  where
+  Limit         :: "i\<Rightarrow>o"  where
     "Limit(i)    \<equiv> Ord(i) \<and> 0<i \<and> (\<forall>y. y<i \<longrightarrow> succ(y)<i)"
 
 abbreviation
@@ -54,7 +54,7 @@
 by (unfold Transset_def, blast)
 
 lemma Transset_Pair_D:
-    "\<lbrakk>Transset(C); <a,b>\<in>C\<rbrakk> \<Longrightarrow> a\<in>C \<and> b\<in>C"
+    "\<lbrakk>Transset(C); \<langle>a,b\<rangle>\<in>C\<rbrakk> \<Longrightarrow> a\<in>C \<and> b\<in>C"
 apply (simp add: Pair_def)
 apply (blast dest: Transset_doubleton_D)
 done
@@ -230,7 +230,7 @@
 done
 
 
-text\<open>Recall that  \<^term>\<open>i \<le> j\<close>  abbreviates  \<^term>\<open>i<succ(j)\<close> \<And>\<close>
+text\<open>Recall that  \<^term>\<open>i \<le> j\<close>  abbreviates  \<^term>\<open>i<succ(j)\<close>!\<close>
 
 lemma le_iff: "i \<le> j <-> i<j | (i=j \<and> Ord(j))"
 by (unfold lt_def, blast)
@@ -267,14 +267,14 @@
 subsection\<open>Natural Deduction Rules for Memrel\<close>
 
 (*The lemmas MemrelI/E give better speed than [iff] here*)
-lemma Memrel_iff [simp]: "<a,b> \<in> Memrel(A) <-> a\<in>b \<and> a\<in>A \<and> b\<in>A"
+lemma Memrel_iff [simp]: "\<langle>a,b\<rangle> \<in> Memrel(A) <-> a\<in>b \<and> a\<in>A \<and> b\<in>A"
 by (unfold Memrel_def, blast)
 
-lemma MemrelI [intro!]: "\<lbrakk>a \<in> b;  a \<in> A;  b \<in> A\<rbrakk> \<Longrightarrow> <a,b> \<in> Memrel(A)"
+lemma MemrelI [intro!]: "\<lbrakk>a \<in> b;  a \<in> A;  b \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Memrel(A)"
 by auto
 
 lemma MemrelE [elim!]:
-    "\<lbrakk><a,b> \<in> Memrel(A);
+    "\<lbrakk>\<langle>a,b\<rangle> \<in> Memrel(A);
         \<lbrakk>a \<in> A;  b \<in> A;  a\<in>b\<rbrakk>  \<Longrightarrow> P\<rbrakk>
      \<Longrightarrow> P"
 by auto
@@ -313,7 +313,7 @@
 
 (*If Transset(A) then Memrel(A) internalizes the membership relation below A*)
 lemma Transset_Memrel_iff:
-    "Transset(A) \<Longrightarrow> <a,b> \<in> Memrel(A) <-> a\<in>b \<and> b\<in>A"
+    "Transset(A) \<Longrightarrow> \<langle>a,b\<rangle> \<in> Memrel(A) <-> a\<in>b \<and> b\<in>A"
 by (unfold Transset_def, blast)
 
 
--- a/src/ZF/Perm.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Perm.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -14,28 +14,28 @@
 
 definition
   (*composition of relations and functions; NOT Suppes's relative product*)
-  comp     :: "[i,i]=>i"      (infixr \<open>O\<close> 60)  where
+  comp     :: "[i,i]\<Rightarrow>i"      (infixr \<open>O\<close> 60)  where
     "r O s \<equiv> {xz \<in> domain(s)*range(r) .
-               \<exists>x y z. xz=<x,z> \<and> <x,y>:s \<and> <y,z>:r}"
+               \<exists>x y z. xz=\<langle>x,z\<rangle> \<and> \<langle>x,y\<rangle>:s \<and> \<langle>y,z\<rangle>:r}"
 
 definition
   (*the identity function for A*)
-  id    :: "i=>i"  where
+  id    :: "i\<Rightarrow>i"  where
     "id(A) \<equiv> (\<lambda>x\<in>A. x)"
 
 definition
   (*one-to-one functions from A to B*)
-  inj   :: "[i,i]=>i"  where
+  inj   :: "[i,i]\<Rightarrow>i"  where
     "inj(A,B) \<equiv> { f \<in> A->B. \<forall>w\<in>A. \<forall>x\<in>A. f`w=f`x \<longrightarrow> w=x}"
 
 definition
   (*onto functions from A to B*)
-  surj  :: "[i,i]=>i"  where
+  surj  :: "[i,i]\<Rightarrow>i"  where
     "surj(A,B) \<equiv> { f \<in> A->B . \<forall>y\<in>B. \<exists>x\<in>A. f`x=y}"
 
 definition
   (*one-to-one and onto functions*)
-  bij   :: "[i,i]=>i"  where
+  bij   :: "[i,i]\<Rightarrow>i"  where
     "bij(A,B) \<equiv> inj(A,B) \<inter> surj(A,B)"
 
 
@@ -89,7 +89,7 @@
 
 text\<open>Good for dealing with sets of pairs, but a bit ugly in use [used in AC]\<close>
 lemma inj_equality:
-    "\<lbrakk><a,b>:f;  <c,b>:f;  f \<in> inj(A,B)\<rbrakk> \<Longrightarrow> a=c"
+    "\<lbrakk>\<langle>a,b\<rangle>:f;  \<langle>c,b\<rangle>:f;  f \<in> inj(A,B)\<rbrakk> \<Longrightarrow> a=c"
 apply (unfold inj_def)
 apply (blast dest: Pair_mem_PiD)
 done
@@ -146,12 +146,12 @@
 
 subsection\<open>Identity Function\<close>
 
-lemma idI [intro!]: "a \<in> A \<Longrightarrow> <a,a> \<in> id(A)"
+lemma idI [intro!]: "a \<in> A \<Longrightarrow> \<langle>a,a\<rangle> \<in> id(A)"
 apply (unfold id_def)
 apply (erule lamI)
 done
 
-lemma idE [elim!]: "\<lbrakk>p \<in> id(A);  \<And>x.\<lbrakk>x \<in> A; p=<x,x>\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow>  P"
+lemma idE [elim!]: "\<lbrakk>p \<in> id(A);  \<And>x.\<lbrakk>x \<in> A; p=\<langle>x,x\<rangle>\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow>  P"
 by (simp add: id_def lam_def, blast)
 
 lemma id_type: "id(A) \<in> A->A"
@@ -192,7 +192,7 @@
 done
 
 text\<open>\<^term>\<open>id\<close> as the identity relation\<close>
-lemma id_iff [simp]: "<x,y> \<in> id(A) \<longleftrightarrow> x=y \<and> y \<in> A"
+lemma id_iff [simp]: "\<langle>x,y\<rangle> \<in> id(A) \<longleftrightarrow> x=y \<and> y \<in> A"
 by auto
 
 
@@ -261,18 +261,18 @@
 
 text\<open>The inductive definition package could derive these theorems for \<^term>\<open>r O s\<close>\<close>
 
-lemma compI [intro]: "\<lbrakk><a,b>:s; <b,c>:r\<rbrakk> \<Longrightarrow> <a,c> \<in> r O s"
+lemma compI [intro]: "\<lbrakk>\<langle>a,b\<rangle>:s; \<langle>b,c\<rangle>:r\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle> \<in> r O s"
 by (unfold comp_def, blast)
 
 lemma compE [elim!]:
     "\<lbrakk>xz \<in> r O s;
-        \<And>x y z. \<lbrakk>xz=<x,z>;  <x,y>:s;  <y,z>:r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+        \<And>x y z. \<lbrakk>xz=\<langle>x,z\<rangle>;  \<langle>x,y\<rangle>:s;  \<langle>y,z\<rangle>:r\<rbrakk> \<Longrightarrow> P\<rbrakk>
      \<Longrightarrow> P"
 by (unfold comp_def, blast)
 
 lemma compEpair:
-    "\<lbrakk><a,c> \<in> r O s;
-        \<And>y. \<lbrakk><a,y>:s;  <y,c>:r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+    "\<lbrakk>\<langle>a,c\<rangle> \<in> r O s;
+        \<And>y. \<lbrakk>\<langle>a,y\<rangle>:s;  \<langle>y,c\<rangle>:r\<rbrakk> \<Longrightarrow> P\<rbrakk>
      \<Longrightarrow> P"
 by (erule compE, simp)
 
@@ -366,7 +366,7 @@
      "\<lbrakk>g \<in> inj(A,B);  f \<in> inj(B,C)\<rbrakk> \<Longrightarrow> (f O g) \<in> inj(A,C)"
 apply (frule inj_is_fun [of g])
 apply (frule inj_is_fun [of f])
-apply (rule_tac d = "%y. converse (g) ` (converse (f) ` y)" in f_imp_injective)
+apply (rule_tac d = "\<lambda>y. converse (g) ` (converse (f) ` y)" in f_imp_injective)
  apply (blast intro: comp_fun, simp)
 done
 
@@ -443,7 +443,7 @@
 lemma comp_eq_id_iff:
     "\<lbrakk>f \<in> A->B;  g \<in> B->A\<rbrakk> \<Longrightarrow> f O g = id(B) \<longleftrightarrow> (\<forall>y\<in>B. f`(g`y)=y)"
 apply (unfold id_def, safe)
- apply (drule_tac t = "%h. h`y " in subst_context)
+ apply (drule_tac t = "\<lambda>h. h`y " in subst_context)
  apply simp
 apply (rule fun_extension)
   apply (blast intro: comp_fun lam_type)
@@ -473,7 +473,7 @@
 lemma inj_disjoint_Un:
      "\<lbrakk>f \<in> inj(A,B);  g \<in> inj(C,D);  B \<inter> D = 0\<rbrakk>
       \<Longrightarrow> (\<lambda>a\<in>A \<union> C. if a \<in> A then f`a else g`a) \<in> inj(A \<union> C, B \<union> D)"
-apply (rule_tac d = "%z. if z \<in> B then converse (f) `z else converse (g) `z"
+apply (rule_tac d = "\<lambda>z. if z \<in> B then converse (f) `z else converse (g) `z"
        in lam_injective)
 apply (auto simp add: inj_is_fun [THEN apply_type])
 done
@@ -546,7 +546,7 @@
 
 lemma inj_extend:
     "\<lbrakk>f \<in> inj(A,B);  a\<notin>A;  b\<notin>B\<rbrakk>
-     \<Longrightarrow> cons(<a,b>,f) \<in> inj(cons(a,A), cons(b,B))"
+     \<Longrightarrow> cons(\<langle>a,b\<rangle>,f) \<in> inj(cons(a,A), cons(b,B))"
 apply (unfold inj_def)
 apply (force intro: apply_type  simp add: fun_extend)
 done
--- a/src/ZF/QPair.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/QPair.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -21,53 +21,53 @@
 \<close>
 
 definition
-  QPair     :: "[i, i] => i"                      (\<open><(_;/ _)>\<close>)  where
+  QPair     :: "[i, i] \<Rightarrow> i"                      (\<open><(_;/ _)>\<close>)  where
     "<a;b> \<equiv> a+b"
 
 definition
-  qfst :: "i => i"  where
+  qfst :: "i \<Rightarrow> i"  where
     "qfst(p) \<equiv> THE a. \<exists>b. p=<a;b>"
 
 definition
-  qsnd :: "i => i"  where
+  qsnd :: "i \<Rightarrow> i"  where
     "qsnd(p) \<equiv> THE b. \<exists>a. p=<a;b>"
 
 definition
-  qsplit    :: "[[i, i] => 'a, i] => 'a::{}"  (*for pattern-matching*)  where
+  qsplit    :: "[[i, i] \<Rightarrow> 'a, i] \<Rightarrow> 'a::{}"  (*for pattern-matching*)  where
     "qsplit(c,p) \<equiv> c(qfst(p), qsnd(p))"
 
 definition
-  qconverse :: "i => i"  where
+  qconverse :: "i \<Rightarrow> i"  where
     "qconverse(r) \<equiv> {z. w \<in> r, \<exists>x y. w=<x;y> \<and> z=<y;x>}"
 
 definition
-  QSigma    :: "[i, i => i] => i"  where
+  QSigma    :: "[i, i \<Rightarrow> i] \<Rightarrow> i"  where
     "QSigma(A,B)  \<equiv>  \<Union>x\<in>A. \<Union>y\<in>B(x). {<x;y>}"
 
 syntax
-  "_QSUM"   :: "[idt, i, i] => i"               (\<open>(3QSUM _ \<in> _./ _)\<close> 10)
+  "_QSUM"   :: "[idt, i, i] \<Rightarrow> i"               (\<open>(3QSUM _ \<in> _./ _)\<close> 10)
 translations
-  "QSUM x \<in> A. B"  => "CONST QSigma(A, %x. B)"
+  "QSUM x \<in> A. B" => "CONST QSigma(A, \<lambda>x. B)"
 
 abbreviation
   qprod  (infixr \<open><*>\<close> 80) where
-  "A <*> B \<equiv> QSigma(A, %_. B)"
+  "A <*> B \<equiv> QSigma(A, \<lambda>_. B)"
 
 definition
-  qsum    :: "[i,i]=>i"                         (infixr \<open><+>\<close> 65)  where
+  qsum    :: "[i,i]\<Rightarrow>i"                         (infixr \<open><+>\<close> 65)  where
     "A <+> B      \<equiv> ({0} <*> A) \<union> ({1} <*> B)"
 
 definition
-  QInl :: "i=>i"  where
+  QInl :: "i\<Rightarrow>i"  where
     "QInl(a)      \<equiv> <0;a>"
 
 definition
-  QInr :: "i=>i"  where
+  QInr :: "i\<Rightarrow>i"  where
     "QInr(b)      \<equiv> <1;b>"
 
 definition
-  qcase     :: "[i=>i, i=>i, i]=>i"  where
-    "qcase(c,d)   \<equiv> qsplit(%y z. cond(y, d(z), c(z)))"
+  qcase     :: "[i\<Rightarrow>i, i\<Rightarrow>i, i]\<Rightarrow>i"  where
+    "qcase(c,d)   \<equiv> qsplit(\<lambda>y z. cond(y, d(z), c(z)))"
 
 
 subsection\<open>Quine ordered pairing\<close>
@@ -149,14 +149,14 @@
 subsubsection\<open>Eliminator: qsplit\<close>
 
 (*A META-equality, so that it applies to higher types as well...*)
-lemma qsplit [simp]: "qsplit(%x y. c(x,y), <a;b>) \<equiv> c(a,b)"
+lemma qsplit [simp]: "qsplit(\<lambda>x y. c(x,y), <a;b>) \<equiv> c(a,b)"
 by (simp add: qsplit_def)
 
 
 lemma qsplit_type [elim!]:
     "\<lbrakk>p \<in> QSigma(A,B);
          \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(<x;y>)
-\<rbrakk> \<Longrightarrow> qsplit(%x y. c(x,y), p) \<in> C(p)"
+\<rbrakk> \<Longrightarrow> qsplit(\<lambda>x y. c(x,y), p) \<in> C(p)"
 by auto
 
 lemma expand_qsplit:
@@ -298,7 +298,7 @@
 lemma Part_QInr: "Part(A <+> B,QInr) = {QInr(y). y \<in> B}"
 by blast
 
-lemma Part_QInr2: "Part(A <+> B, %x. QInr(h(x))) = {QInr(y). y \<in> Part(B,h)}"
+lemma Part_QInr2: "Part(A <+> B, \<lambda>x. QInr(h(x))) = {QInr(y). y \<in> Part(B,h)}"
 by blast
 
 lemma Part_qsum_equality: "C \<subseteq> A <+> B \<Longrightarrow> Part(C,QInl) \<union> Part(C,QInr) = C"
--- a/src/ZF/QUniv.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/QUniv.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -20,7 +20,7 @@
   case_eqns     qcase_QInl qcase_QInr
 
 definition
-  quniv :: "i => i"  where
+  quniv :: "i \<Rightarrow> i"  where
    "quniv(A) \<equiv> Pow(univ(eclose(A)))"
 
 
--- a/src/ZF/Resid/Confluence.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Resid/Confluence.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -6,14 +6,14 @@
 theory Confluence imports Reduction begin
 
 definition
-  confluence    :: "i=>o"  where
+  confluence    :: "i\<Rightarrow>o"  where
     "confluence(R) \<equiv>   
-       \<forall>x y. <x,y> \<in> R \<longrightarrow> (\<forall>z.<x,z> \<in> R \<longrightarrow> (\<exists>u.<y,u> \<in> R \<and> <z,u> \<in> R))"
+       \<forall>x y. \<langle>x,y\<rangle> \<in> R \<longrightarrow> (\<forall>z.\<langle>x,z\<rangle> \<in> R \<longrightarrow> (\<exists>u.\<langle>y,u\<rangle> \<in> R \<and> \<langle>z,u\<rangle> \<in> R))"
 
 definition
   strip         :: "o"  where
     "strip \<equiv> \<forall>x y. (x =\<Longrightarrow> y) \<longrightarrow> 
-                    (\<forall>z.(x =1=> z) \<longrightarrow> (\<exists>u.(y =1=> u) \<and> (z=\<Longrightarrow>u)))" 
+                    (\<forall>z.(x =1\<Rightarrow> z) \<longrightarrow> (\<exists>u.(y =1\<Rightarrow> u) \<and> (z=\<Longrightarrow>u)))" 
 
 
 (* ------------------------------------------------------------------------- *)
@@ -69,11 +69,11 @@
 
 abbreviation
   Sconv1_rel (infixl \<open><-1->\<close> 50) where
-  "a<-1->b \<equiv> <a,b> \<in> Sconv1"
+  "a<-1->b \<equiv> \<langle>a,b\<rangle> \<in> Sconv1"
 
 abbreviation
   Sconv_rel (infixl \<open><-\<longrightarrow>\<close> 50) where
-  "a<-\<longrightarrow>b \<equiv> <a,b> \<in> Sconv"
+  "a<-\<longrightarrow>b \<equiv> \<langle>a,b\<rangle> \<in> Sconv"
   
 inductive
   domains       "Sconv1" \<subseteq> "lambda*lambda"
--- a/src/ZF/Resid/Redex.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Resid/Redex.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -19,28 +19,28 @@
 
 abbreviation
   Ssub_rel  (infixl \<open>\<Longleftarrow>\<close> 70) where
-  "a \<Longleftarrow> b \<equiv> <a,b> \<in> Ssub"
+  "a \<Longleftarrow> b \<equiv> \<langle>a,b\<rangle> \<in> Ssub"
 
 abbreviation
   Scomp_rel  (infixl \<open>\<sim>\<close> 70) where
-  "a \<sim> b \<equiv> <a,b> \<in> Scomp"
+  "a \<sim> b \<equiv> \<langle>a,b\<rangle> \<in> Scomp"
 
 abbreviation
   "regular(a) \<equiv> a \<in> Sreg"
 
-consts union_aux        :: "i=>i"
+consts union_aux        :: "i\<Rightarrow>i"
 primrec (*explicit lambda is required because both arguments of "\<squnion>" vary*)
   "union_aux(Var(n)) =
-     (\<lambda>t \<in> redexes. redexes_case(%j. Var(n), %x. 0, %b x y.0, t))"
+     (\<lambda>t \<in> redexes. redexes_case(\<lambda>j. Var(n), \<lambda>x. 0, \<lambda>b x y.0, t))"
 
   "union_aux(Fun(u)) =
-     (\<lambda>t \<in> redexes. redexes_case(%j. 0, %y. Fun(union_aux(u)`y),
-                                  %b y z. 0, t))"
+     (\<lambda>t \<in> redexes. redexes_case(\<lambda>j. 0, \<lambda>y. Fun(union_aux(u)`y),
+                                  \<lambda>b y z. 0, t))"
 
   "union_aux(App(b,f,a)) =
      (\<lambda>t \<in> redexes.
-        redexes_case(%j. 0, %y. 0,
-                     %c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))"
+        redexes_case(\<lambda>j. 0, \<lambda>y. 0,
+                     \<lambda>c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))"
 
 definition
   union  (infixl \<open>\<squnion>\<close> 70) where
--- a/src/ZF/Resid/Reduction.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Resid/Reduction.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,10 +9,10 @@
 
 consts
   lambda        :: "i"
-  unmark        :: "i=>i"
+  unmark        :: "i\<Rightarrow>i"
 
 abbreviation
-  Apl :: "[i,i]=>i" where
+  Apl :: "[i,i]\<Rightarrow>i" where
   "Apl(n,m) \<equiv> App(0,n,m)"
   
 inductive
@@ -75,19 +75,19 @@
 
 abbreviation
   Sred1_rel (infixl \<open>-1->\<close> 50) where
-  "a -1-> b \<equiv> <a,b> \<in> Sred1"
+  "a -1-> b \<equiv> \<langle>a,b\<rangle> \<in> Sred1"
 
 abbreviation
   Sred_rel (infixl \<open>-\<longrightarrow>\<close> 50) where
-  "a -\<longrightarrow> b \<equiv> <a,b> \<in> Sred"
+  "a -\<longrightarrow> b \<equiv> \<langle>a,b\<rangle> \<in> Sred"
 
 abbreviation
-  Spar_red1_rel (infixl \<open>=1=>\<close> 50) where
-  "a =1=> b \<equiv> <a,b> \<in> Spar_red1"
+  Spar_red1_rel (infixl \<open>=1\<Rightarrow>\<close> 50) where
+  "a =1\<Rightarrow> b \<equiv> \<langle>a,b\<rangle> \<in> Spar_red1"
 
 abbreviation
   Spar_red_rel (infixl \<open>=\<Longrightarrow>\<close> 50) where
-  "a =\<Longrightarrow> b \<equiv> <a,b> \<in> Spar_red"
+  "a =\<Longrightarrow> b \<equiv> \<langle>a,b\<rangle> \<in> Spar_red"
   
   
 inductive
@@ -115,10 +115,10 @@
 inductive
   domains       "Spar_red1" \<subseteq> "lambda*lambda"
   intros
-    beta:       "\<lbrakk>m =1=> m'; n =1=> n'\<rbrakk> \<Longrightarrow> Apl(Fun(m),n) =1=> n'/m'"
-    rvar:       "n \<in> nat \<Longrightarrow> Var(n) =1=> Var(n)"
-    rfun:       "m =1=> m' \<Longrightarrow> Fun(m) =1=> Fun(m')"
-    rapl:       "\<lbrakk>m =1=> m'; n =1=> n'\<rbrakk> \<Longrightarrow> Apl(m,n) =1=> Apl(m',n')"
+    beta:       "\<lbrakk>m =1\<Rightarrow> m'; n =1\<Rightarrow> n'\<rbrakk> \<Longrightarrow> Apl(Fun(m),n) =1\<Rightarrow> n'/m'"
+    rvar:       "n \<in> nat \<Longrightarrow> Var(n) =1\<Rightarrow> Var(n)"
+    rfun:       "m =1\<Rightarrow> m' \<Longrightarrow> Fun(m) =1\<Rightarrow> Fun(m')"
+    rapl:       "\<lbrakk>m =1\<Rightarrow> m'; n =1\<Rightarrow> n'\<rbrakk> \<Longrightarrow> Apl(m,n) =1\<Rightarrow> Apl(m',n')"
   type_intros    red_typechecks
 
 declare Spar_red1.intros [intro, simp]
@@ -126,7 +126,7 @@
 inductive
   domains "Spar_red" \<subseteq> "lambda*lambda"
   intros
-    one_step:   "m =1=> n \<Longrightarrow> m =\<Longrightarrow> n"
+    one_step:   "m =1\<Rightarrow> n \<Longrightarrow> m =\<Longrightarrow> n"
     trans:      "\<lbrakk>m=\<Longrightarrow>n; n=\<Longrightarrow>p\<rbrakk> \<Longrightarrow> m=\<Longrightarrow>p"
   type_intros    Spar_red1.dom_subset [THEN subsetD] red_typechecks
 
@@ -150,7 +150,7 @@
 
 declare bool_typechecks [intro]
 
-inductive_cases  [elim!]: "Fun(t) =1=> Fun(u)"
+inductive_cases  [elim!]: "Fun(t) =1\<Rightarrow> Fun(u)"
 
 
 
@@ -190,10 +190,10 @@
 (* ------------------------------------------------------------------------- *)
 
 
-lemma refl_par_red1: "m \<in> lambda\<Longrightarrow> m =1=> m"
+lemma refl_par_red1: "m \<in> lambda\<Longrightarrow> m =1\<Rightarrow> m"
 by (erule lambda.induct, simp_all)
 
-lemma red1_par_red1: "m-1->n \<Longrightarrow> m=1=>n"
+lemma red1_par_red1: "m-1->n \<Longrightarrow> m=1\<Rightarrow>n"
 by (erule Sred1.induct, simp_all add: refl_par_red1)
 
 lemma red_par_red: "m-\<longrightarrow>n \<Longrightarrow> m=\<Longrightarrow>n"
@@ -214,7 +214,7 @@
 (*      Simulation                                                           *)
 (* ------------------------------------------------------------------------- *)
 
-lemma simulation: "m=1=>n \<Longrightarrow> \<exists>v. m|>v = n \<and> m \<sim> v \<and> regular(v)"
+lemma simulation: "m=1\<Rightarrow>n \<Longrightarrow> \<exists>v. m|>v = n \<and> m \<sim> v \<and> regular(v)"
 by (erule Spar_red1.induct, force+)
 
 
@@ -237,12 +237,12 @@
 (* ------------------------------------------------------------------------- *)
 
 lemma completeness_l [rule_format]:
-     "u \<sim> v \<Longrightarrow> regular(v) \<longrightarrow> unmark(u) =1=> unmark(u|>v)"
+     "u \<sim> v \<Longrightarrow> regular(v) \<longrightarrow> unmark(u) =1\<Rightarrow> unmark(u|>v)"
 apply (erule Scomp.induct)
 apply (auto simp add: unmmark_subst_rec)
 done
 
-lemma completeness: "\<lbrakk>u \<in> lambda; u \<sim> v; regular(v)\<rbrakk> \<Longrightarrow> u =1=> unmark(u|>v)"
+lemma completeness: "\<lbrakk>u \<in> lambda; u \<sim> v; regular(v)\<rbrakk> \<Longrightarrow> u =1\<Rightarrow> unmark(u|>v)"
 by (drule completeness_l, simp_all add: lambda_unmark)
 
 end
--- a/src/ZF/Resid/Residuals.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Resid/Residuals.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -26,7 +26,7 @@
   type_intros    subst_type nat_typechecks redexes.intros bool_typechecks
 
 definition
-  res_func      :: "[i,i]=>i"     (infixl \<open>|>\<close> 70)  where
+  res_func      :: "[i,i]\<Rightarrow>i"     (infixl \<open>|>\<close> 70)  where
   "u |> v \<equiv> THE w. residuals(u,v,w)"
 
 
--- a/src/ZF/Resid/Substitution.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Resid/Substitution.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,7 +8,7 @@
     in the recursive calls ***)
 
 consts
-  lift_aux      :: "i=>i"
+  lift_aux      :: "i\<Rightarrow>i"
 primrec
   "lift_aux(Var(i)) = (\<lambda>k \<in> nat. if i<k then Var(i) else Var(succ(i)))"
 
@@ -17,16 +17,16 @@
   "lift_aux(App(b,f,a)) = (\<lambda>k \<in> nat. App(b, lift_aux(f)`k, lift_aux(a)`k))"
 
 definition
-  lift_rec      :: "[i,i]=> i"  where
+  lift_rec      :: "[i,i]\<Rightarrow> i"  where
     "lift_rec(r,k) \<equiv> lift_aux(r)`k"
 
 abbreviation
-  lift :: "i=>i" where
+  lift :: "i\<Rightarrow>i" where
   "lift(r) \<equiv> lift_rec(r,0)"
 
 
 consts
-  subst_aux     :: "i=>i"
+  subst_aux     :: "i\<Rightarrow>i"
 primrec
   "subst_aux(Var(i)) =
      (\<lambda>r \<in> redexes. \<lambda>k \<in> nat. if k<i then Var(i #- 1)
@@ -38,11 +38,11 @@
      (\<lambda>r \<in> redexes. \<lambda>k \<in> nat. App(b, subst_aux(f)`r`k, subst_aux(a)`r`k))"
 
 definition
-  subst_rec     :: "[i,i,i]=> i"        (**NOTE THE ARGUMENT ORDER BELOW**)  where
+  subst_rec     :: "[i,i,i]\<Rightarrow> i"        (**NOTE THE ARGUMENT ORDER BELOW**)  where
     "subst_rec(u,r,k) \<equiv> subst_aux(r)`u`k"
 
 abbreviation
-  subst :: "[i,i]=>i"  (infixl \<open>'/\<close> 70)  where
+  subst :: "[i,i]\<Rightarrow>i"  (infixl \<open>'/\<close> 70)  where
   "u/v \<equiv> subst_rec(u,v,0)"
 
 
--- a/src/ZF/Sum.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Sum.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,20 +9,20 @@
 
 text\<open>And the "Part" primitive for simultaneous recursive type definitions\<close>
 
-definition sum :: "[i,i]=>i" (infixr \<open>+\<close> 65) where
+definition sum :: "[i,i]\<Rightarrow>i" (infixr \<open>+\<close> 65) where
      "A+B \<equiv> {0}*A \<union> {1}*B"
 
-definition Inl :: "i=>i" where
-     "Inl(a) \<equiv> <0,a>"
+definition Inl :: "i\<Rightarrow>i" where
+     "Inl(a) \<equiv> \<langle>0,a\<rangle>"
 
-definition Inr :: "i=>i" where
-     "Inr(b) \<equiv> <1,b>"
+definition Inr :: "i\<Rightarrow>i" where
+     "Inr(b) \<equiv> \<langle>1,b\<rangle>"
 
-definition "case" :: "[i=>i, i=>i, i]=>i" where
-     "case(c,d) \<equiv> (%<y,z>. cond(y, d(z), c(z)))"
+definition "case" :: "[i\<Rightarrow>i, i\<Rightarrow>i, i]\<Rightarrow>i" where
+     "case(c,d) \<equiv> (\<lambda>\<langle>y,z\<rangle>. cond(y, d(z), c(z)))"
 
   (*operator for selecting out the various summands*)
-definition Part :: "[i,i=>i] => i" where
+definition Part :: "[i,i\<Rightarrow>i] \<Rightarrow> i" where
      "Part(A,h) \<equiv> {x \<in> A. \<exists>z. x = h(z)}"
 
 subsection\<open>Rules for the \<^term>\<open>Part\<close> Primitive\<close>
@@ -154,8 +154,8 @@
 by auto
 
 lemma case_case: "z \<in> A+B \<Longrightarrow>
-        case(c, d, case(%x. Inl(c'(x)), %y. Inr(d'(y)), z)) =
-        case(%x. c(c'(x)), %y. d(d'(y)), z)"
+        case(c, d, case(\<lambda>x. Inl(c'(x)), \<lambda>y. Inr(d'(y)), z)) =
+        case(\<lambda>x. c(c'(x)), \<lambda>y. d(d'(y)), z)"
 by auto
 
 
@@ -179,10 +179,10 @@
 lemma PartD1: "a \<in> Part(A,h) \<Longrightarrow> a \<in> A"
 by (simp add: Part_def)
 
-lemma Part_id: "Part(A,%x. x) = A"
+lemma Part_id: "Part(A,\<lambda>x. x) = A"
 by blast
 
-lemma Part_Inr2: "Part(A+B, %x. Inr(h(x))) = {Inr(y). y \<in> Part(B,h)}"
+lemma Part_Inr2: "Part(A+B, \<lambda>x. Inr(h(x))) = {Inr(y). y \<in> Part(B,h)}"
 by blast
 
 lemma Part_sum_equality: "C \<subseteq> A+B \<Longrightarrow> Part(C,Inl) \<union> Part(C,Inr) = C"
--- a/src/ZF/Trancl.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Trancl.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,44 +8,44 @@
 theory Trancl imports Fixedpt Perm begin
 
 definition
-  refl     :: "[i,i]=>o"  where
-    "refl(A,r) \<equiv> (\<forall>x\<in>A. <x,x> \<in> r)"
+  refl     :: "[i,i]\<Rightarrow>o"  where
+    "refl(A,r) \<equiv> (\<forall>x\<in>A. \<langle>x,x\<rangle> \<in> r)"
 
 definition
-  irrefl   :: "[i,i]=>o"  where
-    "irrefl(A,r) \<equiv> \<forall>x\<in>A. <x,x> \<notin> r"
+  irrefl   :: "[i,i]\<Rightarrow>o"  where
+    "irrefl(A,r) \<equiv> \<forall>x\<in>A. \<langle>x,x\<rangle> \<notin> r"
 
 definition
-  sym      :: "i=>o"  where
-    "sym(r) \<equiv> \<forall>x y. <x,y>: r \<longrightarrow> <y,x>: r"
+  sym      :: "i\<Rightarrow>o"  where
+    "sym(r) \<equiv> \<forall>x y. \<langle>x,y\<rangle>: r \<longrightarrow> \<langle>y,x\<rangle>: r"
 
 definition
-  asym     :: "i=>o"  where
-    "asym(r) \<equiv> \<forall>x y. <x,y>:r \<longrightarrow> \<not> <y,x>:r"
+  asym     :: "i\<Rightarrow>o"  where
+    "asym(r) \<equiv> \<forall>x y. \<langle>x,y\<rangle>:r \<longrightarrow> \<not> \<langle>y,x\<rangle>:r"
 
 definition
-  antisym  :: "i=>o"  where
-    "antisym(r) \<equiv> \<forall>x y.<x,y>:r \<longrightarrow> <y,x>:r \<longrightarrow> x=y"
+  antisym  :: "i\<Rightarrow>o"  where
+    "antisym(r) \<equiv> \<forall>x y.\<langle>x,y\<rangle>:r \<longrightarrow> \<langle>y,x\<rangle>:r \<longrightarrow> x=y"
 
 definition
-  trans    :: "i=>o"  where
-    "trans(r) \<equiv> \<forall>x y z. <x,y>: r \<longrightarrow> <y,z>: r \<longrightarrow> <x,z>: r"
+  trans    :: "i\<Rightarrow>o"  where
+    "trans(r) \<equiv> \<forall>x y z. \<langle>x,y\<rangle>: r \<longrightarrow> \<langle>y,z\<rangle>: r \<longrightarrow> \<langle>x,z\<rangle>: r"
 
 definition
-  trans_on :: "[i,i]=>o"  (\<open>trans[_]'(_')\<close>)  where
+  trans_on :: "[i,i]\<Rightarrow>o"  (\<open>trans[_]'(_')\<close>)  where
     "trans[A](r) \<equiv> \<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A.
-                          <x,y>: r \<longrightarrow> <y,z>: r \<longrightarrow> <x,z>: r"
+                          \<langle>x,y\<rangle>: r \<longrightarrow> \<langle>y,z\<rangle>: r \<longrightarrow> \<langle>x,z\<rangle>: r"
 
 definition
-  rtrancl :: "i=>i"  (\<open>(_^*)\<close> [100] 100)  (*refl/transitive closure*)  where
-    "r^* \<equiv> lfp(field(r)*field(r), %s. id(field(r)) \<union> (r O s))"
+  rtrancl :: "i\<Rightarrow>i"  (\<open>(_^*)\<close> [100] 100)  (*refl/transitive closure*)  where
+    "r^* \<equiv> lfp(field(r)*field(r), \<lambda>s. id(field(r)) \<union> (r O s))"
 
 definition
-  trancl  :: "i=>i"  (\<open>(_^+)\<close> [100] 100)  (*transitive closure*)  where
+  trancl  :: "i\<Rightarrow>i"  (\<open>(_^+)\<close> [100] 100)  (*transitive closure*)  where
     "r^+ \<equiv> r O r^*"
 
 definition
-  equiv    :: "[i,i]=>o"  where
+  equiv    :: "[i,i]\<Rightarrow>o"  where
     "equiv(A,r) \<equiv> r \<subseteq> A*A \<and> refl(A,r) \<and> sym(r) \<and> trans(r)"
 
 
@@ -54,37 +54,37 @@
 subsubsection\<open>irreflexivity\<close>
 
 lemma irreflI:
-    "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> <x,x> \<notin> r\<rbrakk> \<Longrightarrow> irrefl(A,r)"
+    "\<lbrakk>\<And>x. x \<in> A \<Longrightarrow> \<langle>x,x\<rangle> \<notin> r\<rbrakk> \<Longrightarrow> irrefl(A,r)"
 by (simp add: irrefl_def)
 
-lemma irreflE: "\<lbrakk>irrefl(A,r);  x \<in> A\<rbrakk> \<Longrightarrow>  <x,x> \<notin> r"
+lemma irreflE: "\<lbrakk>irrefl(A,r);  x \<in> A\<rbrakk> \<Longrightarrow>  \<langle>x,x\<rangle> \<notin> r"
 by (simp add: irrefl_def)
 
 subsubsection\<open>symmetry\<close>
 
 lemma symI:
-     "\<lbrakk>\<And>x y.<x,y>: r \<Longrightarrow> <y,x>: r\<rbrakk> \<Longrightarrow> sym(r)"
+     "\<lbrakk>\<And>x y.\<langle>x,y\<rangle>: r \<Longrightarrow> \<langle>y,x\<rangle>: r\<rbrakk> \<Longrightarrow> sym(r)"
 by (unfold sym_def, blast)
 
-lemma symE: "\<lbrakk>sym(r); <x,y>: r\<rbrakk>  \<Longrightarrow>  <y,x>: r"
+lemma symE: "\<lbrakk>sym(r); \<langle>x,y\<rangle>: r\<rbrakk>  \<Longrightarrow>  \<langle>y,x\<rangle>: r"
 by (unfold sym_def, blast)
 
 subsubsection\<open>antisymmetry\<close>
 
 lemma antisymI:
-     "\<lbrakk>\<And>x y.\<lbrakk><x,y>: r;  <y,x>: r\<rbrakk> \<Longrightarrow> x=y\<rbrakk> \<Longrightarrow> antisym(r)"
+     "\<lbrakk>\<And>x y.\<lbrakk>\<langle>x,y\<rangle>: r;  \<langle>y,x\<rangle>: r\<rbrakk> \<Longrightarrow> x=y\<rbrakk> \<Longrightarrow> antisym(r)"
 by (simp add: antisym_def, blast)
 
-lemma antisymE: "\<lbrakk>antisym(r); <x,y>: r;  <y,x>: r\<rbrakk>  \<Longrightarrow>  x=y"
+lemma antisymE: "\<lbrakk>antisym(r); \<langle>x,y\<rangle>: r;  \<langle>y,x\<rangle>: r\<rbrakk>  \<Longrightarrow>  x=y"
 by (simp add: antisym_def, blast)
 
 subsubsection\<open>transitivity\<close>
 
-lemma transD: "\<lbrakk>trans(r);  <a,b>:r;  <b,c>:r\<rbrakk> \<Longrightarrow> <a,c>:r"
+lemma transD: "\<lbrakk>trans(r);  \<langle>a,b\<rangle>:r;  \<langle>b,c\<rangle>:r\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle>:r"
 by (unfold trans_def, blast)
 
 lemma trans_onD:
-    "\<lbrakk>trans[A](r);  <a,b>:r;  <b,c>:r;  a \<in> A;  b \<in> A;  c \<in> A\<rbrakk> \<Longrightarrow> <a,c>:r"
+    "\<lbrakk>trans[A](r);  \<langle>a,b\<rangle>:r;  \<langle>b,c\<rangle>:r;  a \<in> A;  b \<in> A;  c \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle>:r"
 by (unfold trans_on_def, blast)
 
 lemma trans_imp_trans_on: "trans(r) \<Longrightarrow> trans[A](r)"
@@ -97,7 +97,7 @@
 subsection\<open>Transitive closure of a relation\<close>
 
 lemma rtrancl_bnd_mono:
-     "bnd_mono(field(r)*field(r), %s. id(field(r)) \<union> (r O s))"
+     "bnd_mono(field(r)*field(r), \<lambda>s. id(field(r)) \<union> (r O s))"
 by (rule bnd_monoI, blast+)
 
 lemma rtrancl_mono: "r<=s \<Longrightarrow> r^* \<subseteq> s^*"
@@ -122,19 +122,19 @@
 done
 
 (*Reflexivity of rtrancl*)
-lemma rtrancl_refl: "\<lbrakk>a \<in> field(r)\<rbrakk> \<Longrightarrow> <a,a> \<in> r^*"
+lemma rtrancl_refl: "\<lbrakk>a \<in> field(r)\<rbrakk> \<Longrightarrow> \<langle>a,a\<rangle> \<in> r^*"
 apply (rule rtrancl_unfold [THEN ssubst])
 apply (erule idI [THEN UnI1])
 done
 
 (*Closure under composition with r  *)
-lemma rtrancl_into_rtrancl: "\<lbrakk><a,b> \<in> r^*;  <b,c> \<in> r\<rbrakk> \<Longrightarrow> <a,c> \<in> r^*"
+lemma rtrancl_into_rtrancl: "\<lbrakk>\<langle>a,b\<rangle> \<in> r^*;  \<langle>b,c\<rangle> \<in> r\<rbrakk> \<Longrightarrow> \<langle>a,c\<rangle> \<in> r^*"
 apply (rule rtrancl_unfold [THEN ssubst])
 apply (rule compI [THEN UnI2], assumption, assumption)
 done
 
 (*rtrancl of r contains all pairs in r  *)
-lemma r_into_rtrancl: "<a,b> \<in> r \<Longrightarrow> <a,b> \<in> r^*"
+lemma r_into_rtrancl: "\<langle>a,b\<rangle> \<in> r \<Longrightarrow> \<langle>a,b\<rangle> \<in> r^*"
 by (rule rtrancl_refl [THEN rtrancl_into_rtrancl], blast+)
 
 (*The premise ensures that r consists entirely of pairs*)
@@ -148,22 +148,22 @@
 (** standard induction rule **)
 
 lemma rtrancl_full_induct [case_names initial step, consumes 1]:
-  "\<lbrakk><a,b> \<in> r^*;
-      \<And>x. x \<in> field(r) \<Longrightarrow> P(<x,x>);
-      \<And>x y z.\<lbrakk>P(<x,y>); <x,y>: r^*; <y,z>: r\<rbrakk>  \<Longrightarrow>  P(<x,z>)\<rbrakk>
-   \<Longrightarrow>  P(<a,b>)"
+  "\<lbrakk>\<langle>a,b\<rangle> \<in> r^*;
+      \<And>x. x \<in> field(r) \<Longrightarrow> P(\<langle>x,x\<rangle>);
+      \<And>x y z.\<lbrakk>P(\<langle>x,y\<rangle>); \<langle>x,y\<rangle>: r^*; \<langle>y,z\<rangle>: r\<rbrakk>  \<Longrightarrow>  P(\<langle>x,z\<rangle>)\<rbrakk>
+   \<Longrightarrow>  P(\<langle>a,b\<rangle>)"
 by (erule def_induct [OF rtrancl_def rtrancl_bnd_mono], blast)
 
 (*nice induction rule.
   Tried adding the typing hypotheses y,z \<in> field(r), but these
   caused expensive case splits!*)
 lemma rtrancl_induct [case_names initial step, induct set: rtrancl]:
-  "\<lbrakk><a,b> \<in> r^*;
+  "\<lbrakk>\<langle>a,b\<rangle> \<in> r^*;
       P(a);
-      \<And>y z.\<lbrakk><a,y> \<in> r^*;  <y,z> \<in> r;  P(y)\<rbrakk> \<Longrightarrow> P(z)
+      \<And>y z.\<lbrakk>\<langle>a,y\<rangle> \<in> r^*;  \<langle>y,z\<rangle> \<in> r;  P(y)\<rbrakk> \<Longrightarrow> P(z)
 \<rbrakk> \<Longrightarrow> P(b)"
 (*by induction on this formula*)
-apply (subgoal_tac "\<forall>y. <a,b> = <a,y> \<longrightarrow> P (y) ")
+apply (subgoal_tac "\<forall>y. \<langle>a,b\<rangle> = \<langle>a,y\<rangle> \<longrightarrow> P (y) ")
 (*now solve first subgoal: this formula is sufficient*)
 apply (erule spec [THEN mp], rule refl)
 (*now do the induction*)
@@ -182,10 +182,10 @@
 
 (*elimination of rtrancl -- by induction on a special formula*)
 lemma rtranclE:
-    "\<lbrakk><a,b> \<in> r^*;  (a=b) \<Longrightarrow> P;
-        \<And>y.\<lbrakk><a,y> \<in> r^*;   <y,b> \<in> r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+    "\<lbrakk>\<langle>a,b\<rangle> \<in> r^*;  (a=b) \<Longrightarrow> P;
+        \<And>y.\<lbrakk>\<langle>a,y\<rangle> \<in> r^*;   \<langle>y,b\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P\<rbrakk>
      \<Longrightarrow> P"
-apply (subgoal_tac "a = b | (\<exists>y. <a,y> \<in> r^* \<and> <y,b> \<in> r) ")
+apply (subgoal_tac "a = b | (\<exists>y. \<langle>a,y\<rangle> \<in> r^* \<and> \<langle>y,b\<rangle> \<in> r) ")
 (*see HOL/trancl*)
 apply blast
 apply (erule rtrancl_induct, blast+)
@@ -207,13 +207,13 @@
 
 (** Conversions between trancl and rtrancl **)
 
-lemma trancl_into_rtrancl: "<a,b> \<in> r^+ \<Longrightarrow> <a,b> \<in> r^*"
+lemma trancl_into_rtrancl: "\<langle>a,b\<rangle> \<in> r^+ \<Longrightarrow> \<langle>a,b\<rangle> \<in> r^*"
 apply (unfold trancl_def)
 apply (blast intro: rtrancl_into_rtrancl)
 done
 
 (*r^+ contains all pairs in r  *)
-lemma r_into_trancl: "<a,b> \<in> r \<Longrightarrow> <a,b> \<in> r^+"
+lemma r_into_trancl: "\<langle>a,b\<rangle> \<in> r \<Longrightarrow> \<langle>a,b\<rangle> \<in> r^+"
 apply (unfold trancl_def)
 apply (blast intro!: rtrancl_refl)
 done
@@ -224,12 +224,12 @@
 
 
 (*intro rule by definition: from r^* and r  *)
-lemma rtrancl_into_trancl1: "\<lbrakk><a,b> \<in> r^*;  <b,c> \<in> r\<rbrakk>   \<Longrightarrow>  <a,c> \<in> r^+"
+lemma rtrancl_into_trancl1: "\<lbrakk>\<langle>a,b\<rangle> \<in> r^*;  \<langle>b,c\<rangle> \<in> r\<rbrakk>   \<Longrightarrow>  \<langle>a,c\<rangle> \<in> r^+"
 by (unfold trancl_def, blast)
 
 (*intro rule from r and r^*  *)
 lemma rtrancl_into_trancl2:
-    "\<lbrakk><a,b> \<in> r;  <b,c> \<in> r^*\<rbrakk>   \<Longrightarrow>  <a,c> \<in> r^+"
+    "\<lbrakk>\<langle>a,b\<rangle> \<in> r;  \<langle>b,c\<rangle> \<in> r^*\<rbrakk>   \<Longrightarrow>  \<langle>a,c\<rangle> \<in> r^+"
 apply (erule rtrancl_induct)
  apply (erule r_into_trancl)
 apply (blast intro: r_into_trancl trancl_trans)
@@ -237,14 +237,14 @@
 
 (*Nice induction rule for trancl*)
 lemma trancl_induct [case_names initial step, induct set: trancl]:
-  "\<lbrakk><a,b> \<in> r^+;
-      \<And>y.  \<lbrakk><a,y> \<in> r\<rbrakk> \<Longrightarrow> P(y);
-      \<And>y z.\<lbrakk><a,y> \<in> r^+;  <y,z> \<in> r;  P(y)\<rbrakk> \<Longrightarrow> P(z)
+  "\<lbrakk>\<langle>a,b\<rangle> \<in> r^+;
+      \<And>y.  \<lbrakk>\<langle>a,y\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P(y);
+      \<And>y z.\<lbrakk>\<langle>a,y\<rangle> \<in> r^+;  \<langle>y,z\<rangle> \<in> r;  P(y)\<rbrakk> \<Longrightarrow> P(z)
 \<rbrakk> \<Longrightarrow> P(b)"
 apply (rule compEpair)
 apply (unfold trancl_def, assumption)
 (*by induction on this formula*)
-apply (subgoal_tac "\<forall>z. <y,z> \<in> r \<longrightarrow> P (z) ")
+apply (subgoal_tac "\<forall>z. \<langle>y,z\<rangle> \<in> r \<longrightarrow> P (z) ")
 (*now solve first subgoal: this formula is sufficient*)
  apply blast
 apply (erule rtrancl_induct)
@@ -253,11 +253,11 @@
 
 (*elimination of r^+ -- NOT an induction rule*)
 lemma tranclE:
-    "\<lbrakk><a,b> \<in> r^+;
-        <a,b> \<in> r \<Longrightarrow> P;
-        \<And>y.\<lbrakk><a,y> \<in> r^+; <y,b> \<in> r\<rbrakk> \<Longrightarrow> P
+    "\<lbrakk>\<langle>a,b\<rangle> \<in> r^+;
+        \<langle>a,b\<rangle> \<in> r \<Longrightarrow> P;
+        \<And>y.\<lbrakk>\<langle>a,y\<rangle> \<in> r^+; \<langle>y,b\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P
 \<rbrakk> \<Longrightarrow> P"
-apply (subgoal_tac "<a,b> \<in> r | (\<exists>y. <a,y> \<in> r^+ \<and> <y,b> \<in> r) ")
+apply (subgoal_tac "\<langle>a,b\<rangle> \<in> r | (\<exists>y. \<langle>a,y\<rangle> \<in> r^+ \<and> \<langle>y,b\<rangle> \<in> r) ")
 apply blast
 apply (rule compEpair)
 apply (unfold trancl_def, assumption)
@@ -320,7 +320,7 @@
 
 (** rtrancl **)
 
-lemma rtrancl_converseD: "<x,y>:converse(r)^* \<Longrightarrow> <x,y>:converse(r^*)"
+lemma rtrancl_converseD: "\<langle>x,y\<rangle>:converse(r)^* \<Longrightarrow> \<langle>x,y\<rangle>:converse(r^*)"
 apply (rule converseI)
 apply (frule rtrancl_type [THEN subsetD])
 apply (erule rtrancl_induct)
@@ -328,7 +328,7 @@
 apply (blast intro: r_into_rtrancl rtrancl_trans)
 done
 
-lemma rtrancl_converseI: "<x,y>:converse(r^*) \<Longrightarrow> <x,y>:converse(r)^*"
+lemma rtrancl_converseI: "\<langle>x,y\<rangle>:converse(r^*) \<Longrightarrow> \<langle>x,y\<rangle>:converse(r)^*"
 apply (drule converseD)
 apply (frule rtrancl_type [THEN subsetD])
 apply (erule rtrancl_induct)
@@ -344,12 +344,12 @@
 
 (** trancl **)
 
-lemma trancl_converseD: "<a, b>:converse(r)^+ \<Longrightarrow> <a, b>:converse(r^+)"
+lemma trancl_converseD: "\<langle>a, b\<rangle>:converse(r)^+ \<Longrightarrow> \<langle>a, b\<rangle>:converse(r^+)"
 apply (erule trancl_induct)
 apply (auto intro: r_into_trancl trancl_trans)
 done
 
-lemma trancl_converseI: "<x,y>:converse(r^+) \<Longrightarrow> <x,y>:converse(r)^+"
+lemma trancl_converseI: "\<langle>x,y\<rangle>:converse(r^+) \<Longrightarrow> \<langle>x,y\<rangle>:converse(r)^+"
 apply (drule converseD)
 apply (erule trancl_induct)
 apply (auto intro: r_into_trancl trancl_trans)
@@ -362,8 +362,8 @@
 done
 
 lemma converse_trancl_induct [case_names initial step, consumes 1]:
-"\<lbrakk><a, b>:r^+; \<And>y. <y, b> :r \<Longrightarrow> P(y);
-      \<And>y z. \<lbrakk><y, z> \<in> r; <z, b> \<in> r^+; P(z)\<rbrakk> \<Longrightarrow> P(y)\<rbrakk>
+"\<lbrakk>\<langle>a, b\<rangle>:r^+; \<And>y. \<langle>y, b\<rangle> :r \<Longrightarrow> P(y);
+      \<And>y z. \<lbrakk>\<langle>y, z\<rangle> \<in> r; \<langle>z, b\<rangle> \<in> r^+; P(z)\<rbrakk> \<Longrightarrow> P(y)\<rbrakk>
        \<Longrightarrow> P(a)"
 apply (drule converseI)
 apply (simp (no_asm_use) add: trancl_converse [symmetric])
--- a/src/ZF/UNITY/AllocBase.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/AllocBase.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -20,13 +20,13 @@
   Nclients_pos: "Nclients \<in> nat-{0}"
   
 text\<open>This function merely sums the elements of a list\<close>
-consts tokens :: "i =>i"
+consts tokens :: "i \<Rightarrow>i"
        item :: i (* Items to be merged/distributed *)
 primrec 
   "tokens(Nil) = 0"
   "tokens (Cons(x,xs)) = x #+ tokens(xs)"
 
-consts bag_of :: "i => i"
+consts bag_of :: "i \<Rightarrow> i"
 primrec
   "bag_of(Nil)    = 0"
   "bag_of(Cons(x,xs)) = {#x#} +# bag_of(xs)"
@@ -34,28 +34,28 @@
 
 text\<open>Definitions needed in Client.thy.  We define a recursive predicate
 using 0 and 1 to code the truth values.\<close>
-consts all_distinct0 :: "i=>i"
+consts all_distinct0 :: "i\<Rightarrow>i"
 primrec
   "all_distinct0(Nil) = 1"
   "all_distinct0(Cons(a, l)) =
      (if a \<in> set_of_list(l) then 0 else all_distinct0(l))"
 
 definition
-  all_distinct  :: "i=>o"  where
+  all_distinct  :: "i\<Rightarrow>o"  where
    "all_distinct(l) \<equiv> all_distinct0(l)=1"
   
 definition  
-  state_of :: "i =>i" \<comment> \<open>coersion from anyting to state\<close>  where
+  state_of :: "i \<Rightarrow>i" \<comment> \<open>coersion from anyting to state\<close>  where
    "state_of(s) \<equiv> if s \<in> state then s else st0"
 
 definition
-  lift :: "i =>(i=>i)" \<comment> \<open>simplifies the expression of programs\<close>  where
-   "lift(x) \<equiv> %s. s`x"
+  lift :: "i \<Rightarrow>(i\<Rightarrow>i)" \<comment> \<open>simplifies the expression of programs\<close>  where
+   "lift(x) \<equiv> \<lambda>s. s`x"
 
 text\<open>function to show that the set of variables is infinite\<close>
 consts
-  nat_list_inj :: "i=>i"
-  var_inj      :: "i=>i"
+  nat_list_inj :: "i\<Rightarrow>i"
+  var_inj      :: "i\<Rightarrow>i"
 
 primrec
   "nat_list_inj(0) = Nil"
@@ -65,7 +65,7 @@
   "var_inj(Var(l)) = length(l)"
 
 definition
-  nat_var_inj  :: "i=>i"  where
+  nat_var_inj  :: "i\<Rightarrow>i"  where
   "nat_var_inj(n) \<equiv> Var(nat_list_inj(n))"
 
 
@@ -106,13 +106,13 @@
 by (erule list.induct, auto)
 
 lemma tokens_mono_aux [rule_format]:
-     "xs\<in>list(A) \<Longrightarrow> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)  
+     "xs\<in>list(A) \<Longrightarrow> \<forall>ys\<in>list(A). \<langle>xs, ys\<rangle>\<in>prefix(A)  
    \<longrightarrow> tokens(xs) \<le> tokens(ys)"
 apply (induct_tac "xs")
 apply (auto dest: gen_prefix.dom_subset [THEN subsetD] simp add: prefix_def)
 done
 
-lemma tokens_mono: "<xs, ys>\<in>prefix(A) \<Longrightarrow> tokens(xs) \<le> tokens(ys)"
+lemma tokens_mono: "\<langle>xs, ys\<rangle>\<in>prefix(A) \<Longrightarrow> tokens(xs) \<le> tokens(ys)"
 apply (cut_tac prefix_type)
 apply (blast intro: tokens_mono_aux)
 done
@@ -147,7 +147,7 @@
 done
 
 lemma bag_of_mono_aux [rule_format]:
-     "xs\<in>list(A) \<Longrightarrow> \<forall>ys\<in>list(A). <xs, ys>\<in>prefix(A)  
+     "xs\<in>list(A) \<Longrightarrow> \<forall>ys\<in>list(A). \<langle>xs, ys\<rangle>\<in>prefix(A)  
       \<longrightarrow> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)"
 apply (induct_tac "xs", simp_all, clarify) 
 apply (frule_tac l = ys in bag_of_multiset)
@@ -158,7 +158,7 @@
 done
 
 lemma bag_of_mono [intro]:
-     "\<lbrakk><xs, ys>\<in>prefix(A); xs\<in>list(A); ys\<in>list(A)\<rbrakk>
+     "\<lbrakk>\<langle>xs, ys\<rangle>\<in>prefix(A); xs\<in>list(A); ys\<in>list(A)\<rbrakk>
       \<Longrightarrow> <bag_of(xs), bag_of(ys)>\<in>MultLe(A, r)"
 apply (blast intro: bag_of_mono_aux)
 done
@@ -212,7 +212,7 @@
      "l\<in>list(A) \<Longrightarrow>  
   C \<subseteq> nat \<Longrightarrow>  
   bag_of(sublist(l, C)) =  
-      msetsum(%i. {#nth(i, l)#}, C \<inter> length(l), A)"
+      msetsum(\<lambda>i. {#nth(i, l)#}, C \<inter> length(l), A)"
 apply (erule list_append_induct)
 apply (simp (no_asm))
 apply (simp (no_asm_simp) add: sublist_append nth_append bag_of_sublist_lemma munion_commute bag_of_sublist_lemma msetsum_multiset munion_0)
@@ -227,8 +227,8 @@
 (*eliminating the assumption C<=nat*)
 lemma bag_of_sublist:
      "l\<in>list(A) \<Longrightarrow>  
-  bag_of(sublist(l, C)) = msetsum(%i. {#nth(i, l)#}, C \<inter> length(l), A)"
-apply (subgoal_tac " bag_of (sublist (l, C \<inter> nat)) = msetsum (%i. {#nth (i, l) #}, C \<inter> length (l), A) ")
+  bag_of(sublist(l, C)) = msetsum(\<lambda>i. {#nth(i, l)#}, C \<inter> length(l), A)"
+apply (subgoal_tac " bag_of (sublist (l, C \<inter> nat)) = msetsum (\<lambda>i. {#nth (i, l) #}, C \<inter> length (l), A) ")
 apply (simp add: sublist_Int_eq)
 apply (simp add: bag_of_sublist_lemma2 Int_lower2 Int_assoc nat_Int_length_eq)
 done
@@ -257,7 +257,7 @@
      "\<lbrakk>Finite(I); \<forall>i\<in>I. \<forall>j\<in>I. i\<noteq>j \<longrightarrow> A(i) \<inter> A(j) = 0;  
         l\<in>list(B)\<rbrakk>  
       \<Longrightarrow> bag_of(sublist(l, \<Union>i\<in>I. A(i))) =   
-          (msetsum(%i. bag_of(sublist(l, A(i))), I, B)) "
+          (msetsum(\<lambda>i. bag_of(sublist(l, A(i))), I, B)) "
 apply (simp (no_asm_simp) del: UN_simps
            add: UN_simps [symmetric] bag_of_sublist)
 apply (subst  msetsum_UN_disjoint [of _ _ _ "length (l)"])
@@ -308,7 +308,7 @@
 (** Used in ClientImp **)
 
 lemma gen_Increains_state_of_eq: 
-     "Increasing(A, r, %s. f(state_of(s))) = Increasing(A, r, f)"
+     "Increasing(A, r, \<lambda>s. f(state_of(s))) = Increasing(A, r, f)"
 apply (unfold Increasing_def, auto)
 done
 
@@ -318,7 +318,7 @@
       gen_Increains_state_of_eq [THEN equalityD2, THEN subsetD]
 
 lemma Follows_state_of_eq: 
-     "Follows(A, r, %s. f(state_of(s)), %s. g(state_of(s))) =   
+     "Follows(A, r, \<lambda>s. f(state_of(s)), \<lambda>s. g(state_of(s))) =   
       Follows(A, r, f, g)"
 apply (unfold Follows_def Increasing_def, auto)
 done
@@ -383,7 +383,7 @@
 
 lemma setsum_nsetsum_eq: 
      "\<lbrakk>Finite(A); \<forall>x\<in>A. g(x)\<in>nat\<rbrakk> 
-      \<Longrightarrow> setsum(%x. $#(g(x)), A) = $# nsetsum(%x. g(x), A)"
+      \<Longrightarrow> setsum(\<lambda>x. $#(g(x)), A) = $# nsetsum(\<lambda>x. g(x), A)"
 apply (erule Finite_induct)
 apply (auto simp add: int_of_add)
 done
--- a/src/ZF/UNITY/AllocImpl.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/AllocImpl.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -25,7 +25,7 @@
 
 definition
   "alloc_giv_act \<equiv>
-       {<s, t> \<in> state*state.
+       {\<langle>s, t\<rangle> \<in> state*state.
         \<exists>k. k = length(s`giv) \<and>
             t = s(giv := s`giv @ [nth(k, s`ask)],
                   available_tok := s`available_tok #- nth(k, s`ask)) \<and>
@@ -33,7 +33,7 @@
 
 definition
   "alloc_rel_act \<equiv>
-       {<s, t> \<in> state*state.
+       {\<langle>s, t\<rangle> \<in> state*state.
         t = s(available_tok := s`available_tok #+ nth(s`NbR, s`rel),
               NbR := succ(s`NbR)) \<and>
         s`NbR < length(s`rel)}"
@@ -202,7 +202,7 @@
 apply (auto intro!: AlwaysI simp add: guar_def)
 apply (subgoal_tac "G \<in> preserves (lift (giv))")
  prefer 2 apply (simp add: alloc_prog_ok_iff)
-apply (rule_tac P = "%x y. <x,y> \<in> prefix(tokbag)" and A = "list(nat)"
+apply (rule_tac P = "\<lambda>x y. \<langle>x,y\<rangle> \<in> prefix(tokbag)" and A = "list(nat)"
        in stable_Join_Stable)
 apply safety
  prefer 2 apply (simp add: lift_def, clarify)
@@ -573,7 +573,7 @@
 apply (rule LeadsTo_weaken_R)
 apply (rule Always_LeadsToD [OF safety length_ask_giv], assumption+, clarify)
 apply (simp add: INT_iff)
-apply (drule_tac x = "length(x ` giv)" and P = "%x. f (x) \<le> NbT" for f in bspec)
+apply (drule_tac x = "length(x ` giv)" and P = "\<lambda>x. f (x) \<le> NbT" for f in bspec)
 apply simp
 apply (blast intro: le_trans)
 done
--- a/src/ZF/UNITY/ClientImpl.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/ClientImpl.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -26,7 +26,7 @@
 definition
  (** Release some client_tokens **)
     "client_rel_act \<equiv>
-     {<s,t> \<in> state*state.
+     {\<langle>s,t\<rangle> \<in> state*state.
       \<exists>nrel \<in> nat. nrel = length(s`rel) \<and>
                    t = s(rel:=(s`rel)@[nth(nrel, s`giv)]) \<and>
                    nrel < length(s`giv) \<and>
@@ -36,11 +36,11 @@
   (** Including t=s suppresses fairness, allowing the non-trivial part
       of the action to be ignored **)
 definition
-  "client_tok_act \<equiv> {<s,t> \<in> state*state. t=s |
+  "client_tok_act \<equiv> {\<langle>s,t\<rangle> \<in> state*state. t=s |
                      t = s(tok:=succ(s`tok mod NbT))}"
 
 definition
-  "client_ask_act \<equiv> {<s,t> \<in> state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}"
+  "client_ask_act \<equiv> {\<langle>s,t\<rangle> \<in> state*state. t=s | (t=s(ask:=s`ask@[s`tok]))}"
 
 definition
   "client_prog \<equiv>
@@ -178,15 +178,15 @@
 by (force intro!: AlwaysI client_prog_Join_Stable_rel_le_giv)
 
 lemma def_act_eq:
-     "A \<equiv> {<s, t> \<in> state*state. P(s, t)} \<Longrightarrow> A={<s, t> \<in> state*state. P(s, t)}"
+     "A \<equiv> {\<langle>s, t\<rangle> \<in> state*state. P(s, t)} \<Longrightarrow> A={\<langle>s, t\<rangle> \<in> state*state. P(s, t)}"
 by auto
 
-lemma act_subset: "A={<s,t> \<in> state*state. P(s, t)} \<Longrightarrow> A<=state*state"
+lemma act_subset: "A={\<langle>s,t\<rangle> \<in> state*state. P(s, t)} \<Longrightarrow> A<=state*state"
 by auto
 
 lemma transient_lemma:
 "client_prog \<in>
-  transient({s \<in> state. s`rel = k \<and> <k, h> \<in> strict_prefix(nat)
+  transient({s \<in> state. s`rel = k \<and> \<langle>k, h\<rangle> \<in> strict_prefix(nat)
    \<and> <h, s`giv> \<in> prefix(nat) \<and> h pfixGe s`ask})"
 apply (rule_tac act = client_rel_act in transientI)
 apply (simp (no_asm) add: client_prog_def [THEN def_prg_Acts])
@@ -209,7 +209,7 @@
 done
 
 lemma strict_prefix_is_prefix:
-    "<xs, ys> \<in> strict_prefix(A) \<longleftrightarrow>  <xs, ys> \<in> prefix(A) \<and> xs\<noteq>ys"
+    "\<langle>xs, ys\<rangle> \<in> strict_prefix(A) \<longleftrightarrow>  \<langle>xs, ys\<rangle> \<in> prefix(A) \<and> xs\<noteq>ys"
 apply (unfold strict_prefix_def id_def lam_def)
 apply (auto dest: prefix_type [THEN subsetD])
 done
@@ -217,7 +217,7 @@
 lemma induct_lemma:
 "\<lbrakk>client_prog \<squnion> G \<in> Incr(lift(giv)); client_prog ok G; G \<in> program\<rbrakk>
   \<Longrightarrow> client_prog \<squnion> G \<in>
-  {s \<in> state. s`rel = k \<and> <k,h> \<in> strict_prefix(nat)
+  {s \<in> state. s`rel = k \<and> \<langle>k,h\<rangle> \<in> strict_prefix(nat)
    \<and> <h, s`giv> \<in> prefix(nat) \<and> h pfixGe s`ask}
         \<longmapsto>w {s \<in> state. <k, s`rel> \<in> strict_prefix(nat)
                           \<and> <s`rel, s`giv> \<in> prefix(nat) \<and>
--- a/src/ZF/UNITY/Comp.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Comp.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -18,33 +18,33 @@
 theory Comp imports Union Increasing begin
 
 definition
-  component :: "[i,i]=>o"  (infixl \<open>component\<close> 65)  where
+  component :: "[i,i]\<Rightarrow>o"  (infixl \<open>component\<close> 65)  where
   "F component H \<equiv> (\<exists>G. F \<squnion> G = H)"
 
 definition
-  strict_component :: "[i,i]=>o" (infixl \<open>strict'_component\<close> 65)  where
+  strict_component :: "[i,i]\<Rightarrow>o" (infixl \<open>strict'_component\<close> 65)  where
   "F strict_component H \<equiv> F component H \<and> F\<noteq>H"
 
 definition
   (* A stronger form of the component relation *)
-  component_of :: "[i,i]=>o"   (infixl \<open>component'_of\<close> 65)  where
+  component_of :: "[i,i]\<Rightarrow>o"   (infixl \<open>component'_of\<close> 65)  where
   "F component_of H  \<equiv> (\<exists>G. F ok G \<and> F \<squnion> G = H)"
   
 definition
-  strict_component_of :: "[i,i]=>o" (infixl \<open>strict'_component'_of\<close> 65)  where
+  strict_component_of :: "[i,i]\<Rightarrow>o" (infixl \<open>strict'_component'_of\<close> 65)  where
   "F strict_component_of H  \<equiv> F component_of H  \<and> F\<noteq>H"
 
 definition
   (* Preserves a state function f, in particular a variable *)
-  preserves :: "(i=>i)=>i"  where
+  preserves :: "(i\<Rightarrow>i)\<Rightarrow>i"  where
   "preserves(f) \<equiv> {F:program. \<forall>z. F: stable({s:state. f(s) = z})}"
 
 definition
-  fun_pair :: "[i=>i, i =>i] =>(i=>i)"  where
-  "fun_pair(f, g) \<equiv> %x. <f(x), g(x)>"
+  fun_pair :: "[i\<Rightarrow>i, i \<Rightarrow>i] \<Rightarrow>(i\<Rightarrow>i)"  where
+  "fun_pair(f, g) \<equiv> \<lambda>x. <f(x), g(x)>"
 
 definition
-  localize  :: "[i=>i, i] => i"  where
+  localize  :: "[i\<Rightarrow>i, i] \<Rightarrow> i"  where
   "localize(f,F) \<equiv> mk_program(Init(F), Acts(F),
                        AllowedActs(F) \<inter> (\<Union>G\<in>preserves(f). Acts(G)))"
 
@@ -156,7 +156,7 @@
 done
 
 lemma preserves_imp_eq: 
-     "\<lbrakk>F \<in> preserves(f);  act \<in> Acts(F);  <s,t> \<in> act\<rbrakk> \<Longrightarrow> f(s) = f(t)"
+     "\<lbrakk>F \<in> preserves(f);  act \<in> Acts(F);  \<langle>s,t\<rangle> \<in> act\<rbrakk> \<Longrightarrow> f(s) = f(t)"
 apply (unfold preserves_def stable_def constrains_def)
 apply (subgoal_tac "s \<in> state \<and> t \<in> state")
  prefer 2 apply (blast dest!: Acts_type [THEN subsetD], force) 
@@ -226,11 +226,11 @@
  "\<lbrakk>F \<in> preserves(f); \<forall>x \<in> state. f(x):A\<rbrakk> \<Longrightarrow> F \<in> Increasing.increasing(A, r, f)"
 apply (unfold Increasing.increasing_def)
 apply (auto intro: preserves_into_program)
-apply (rule_tac P = "%x. <k, x>:r" in preserves_imp_stable, auto)
+apply (rule_tac P = "\<lambda>x. \<langle>k, x\<rangle>:r" in preserves_imp_stable, auto)
 done
 
 lemma preserves_id_subset_stable: 
- "st_set(A) \<Longrightarrow> preserves(%x. x) \<subseteq> stable(A)"
+ "st_set(A) \<Longrightarrow> preserves(\<lambda>x. x) \<subseteq> stable(A)"
 apply (unfold preserves_def stable_def constrains_def, auto)
 apply (drule_tac x = xb in spec)
 apply (drule_tac x = act in bspec)
@@ -238,7 +238,7 @@
 done
 
 lemma preserves_id_imp_stable:
-     "\<lbrakk>F \<in> preserves(%x. x); st_set(A)\<rbrakk> \<Longrightarrow> F \<in> stable(A)"
+     "\<lbrakk>F \<in> preserves(\<lambda>x. x); st_set(A)\<rbrakk> \<Longrightarrow> F \<in> stable(A)"
 by (blast intro: preserves_id_subset_stable [THEN subsetD])
 
 (** Added by Sidi **)
--- a/src/ZF/UNITY/Constrains.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Constrains.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,8 +9,8 @@
 imports UNITY
 begin
 
-consts traces :: "[i, i] => i"
-  (* Initial states and program => (final state, reversed trace to it)... 
+consts traces :: "[i, i] \<Rightarrow> i"
+  (* Initial states and program \<Rightarrow> (final state, reversed trace to it)... 
       the domain may also be state*list(state) *)
 inductive 
   domains 
@@ -20,13 +20,13 @@
          (*Initial trace is empty*)
     Init: "s: init \<Longrightarrow> <s,[]> \<in> traces(init,acts)"
 
-    Acts: "\<lbrakk>act:acts;  <s,evs> \<in> traces(init,acts);  <s,s'>: act\<rbrakk>
+    Acts: "\<lbrakk>act:acts;  \<langle>s,evs\<rangle> \<in> traces(init,acts);  <s,s'>: act\<rbrakk>
            \<Longrightarrow> <s', Cons(s,evs)> \<in> traces(init, acts)"
   
   type_intros list.intros UnI1 UnI2 UN_I fieldI2 fieldI1
 
 
-consts reachable :: "i=>i"
+consts reachable :: "i\<Rightarrow>i"
 inductive
   domains
   "reachable(F)" \<subseteq> "Init(F) \<union> (\<Union>act\<in>Acts(F). field(act))"
@@ -40,20 +40,20 @@
 
   
 definition
-  Constrains :: "[i,i] => i"  (infixl \<open>Co\<close> 60)  where
+  Constrains :: "[i,i] \<Rightarrow> i"  (infixl \<open>Co\<close> 60)  where
   "A Co B \<equiv> {F:program. F:(reachable(F) \<inter> A) co B}"
 
 definition
-  op_Unless  :: "[i, i] => i"  (infixl \<open>Unless\<close> 60)  where
+  op_Unless  :: "[i, i] \<Rightarrow> i"  (infixl \<open>Unless\<close> 60)  where
   "A Unless B \<equiv> (A-B) Co (A \<union> B)"
 
 definition
-  Stable     :: "i => i"  where
+  Stable     :: "i \<Rightarrow> i"  where
   "Stable(A) \<equiv> A Co A"
 
 definition
   (*Always is the weak form of "invariant"*)
-  Always :: "i => i"  where
+  Always :: "i \<Rightarrow> i"  where
   "Always(A) \<equiv> initially(A) \<inter> Stable(A)"
 
 
@@ -80,7 +80,7 @@
 declare state_Int_reachable [iff]
 
 lemma reachable_equiv_traces: 
-"F \<in> program \<Longrightarrow> reachable(F)={s \<in> state. \<exists>evs. <s,evs>:traces(Init(F), Acts(F))}"
+"F \<in> program \<Longrightarrow> reachable(F)={s \<in> state. \<exists>evs. \<langle>s,evs\<rangle>:traces(Init(F), Acts(F))}"
 apply (rule equalityI, safe)
 apply (blast dest: reachable_type [THEN subsetD])
 apply (erule_tac [2] traces.induct)
--- a/src/ZF/UNITY/Distributor.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Distributor.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -13,7 +13,7 @@
 text\<open>spec (14)\<close>
 
 definition
-  distr_follows :: "[i, i, i, i =>i] =>i"  where
+  distr_follows :: "[i, i, i, i \<Rightarrow>i] \<Rightarrow>i"  where
     "distr_follows(A, In, iIn, Out) \<equiv>
      (lift(In) IncreasingWrt prefix(A)/list(A)) \<inter>
      (lift(iIn) IncreasingWrt prefix(nat)/list(nat)) \<inter>
@@ -22,17 +22,17 @@
          (\<Inter>n \<in> Nclients.
           lift(Out(n))
               Fols
-          (%s. sublist(s`In, {k \<in> nat. k<length(s`iIn) \<and> nth(k, s`iIn) = n}))
+          (\<lambda>s. sublist(s`In, {k \<in> nat. k<length(s`iIn) \<and> nth(k, s`iIn) = n}))
           Wrt prefix(A)/list(A))"
 
 definition
-  distr_allowed_acts :: "[i=>i]=>i"  where
+  distr_allowed_acts :: "[i\<Rightarrow>i]\<Rightarrow>i"  where
     "distr_allowed_acts(Out) \<equiv>
      {D \<in> program. AllowedActs(D) =
      cons(id(state), \<Union>G \<in> (\<Inter>n\<in>nat. preserves(lift(Out(n)))). Acts(G))}"
 
 definition
-  distr_spec :: "[i, i, i, i =>i]=>i"  where
+  distr_spec :: "[i, i, i, i \<Rightarrow>i]\<Rightarrow>i"  where
     "distr_spec(A, In, iIn, Out) \<equiv>
      distr_follows(A, In, iIn, Out) \<inter> distr_allowed_acts(Out)"
 
@@ -105,7 +105,7 @@
    D \<squnion> G \<in> Always({s \<in> state.
           \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients})\<rbrakk>
   \<Longrightarrow> D \<squnion> G \<in> Always
-          ({s \<in> state. msetsum(%n. bag_of (sublist(s`In,
+          ({s \<in> state. msetsum(\<lambda>n. bag_of (sublist(s`In,
                        {k \<in> nat. k < length(s`iIn) \<and>
                           nth(k, s`iIn)= n})),
                          Nclients, A) =
@@ -131,7 +131,7 @@
 apply (subgoal_tac "length (s ` iIn) \<in> nat")
 apply typecheck
 apply (subgoal_tac "m \<in> nat")
-apply (drule_tac x = "nth(m, s`iIn) " and P = "%elt. X (elt) \<longrightarrow> elt<Nclients" for X in bspec)
+apply (drule_tac x = "nth(m, s`iIn) " and P = "\<lambda>elt. X (elt) \<longrightarrow> elt<Nclients" for X in bspec)
 apply (simp add: ltI)
 apply (simp_all add: Ord_mem_iff_lt)
 apply (blast dest: ltD)
@@ -144,9 +144,9 @@
         Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iIn). elt < Nclients}))
       guarantees
        (\<Inter>n \<in> Nclients.
-        (%s. msetsum(%i. bag_of(s`Out(i)),  Nclients, A))
+        (\<lambda>s. msetsum(\<lambda>i. bag_of(s`Out(i)),  Nclients, A))
         Fols
-        (%s. bag_of(sublist(s`In, length(s`iIn))))
+        (\<lambda>s. bag_of(sublist(s`In, length(s`iIn))))
         Wrt MultLe(A, r)/Mult(A))"
 apply (cut_tac distr_spec)
 apply (rule guaranteesI, clarify)
--- a/src/ZF/UNITY/FP.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/FP.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,11 +12,11 @@
 theory FP imports UNITY begin
 
 definition   
-  FP_Orig :: "i=>i"  where
+  FP_Orig :: "i\<Rightarrow>i"  where
     "FP_Orig(F) \<equiv> \<Union>({A \<in> Pow(state). \<forall>B. F \<in> stable(A \<inter> B)})"
 
 definition
-  FP :: "i=>i"  where
+  FP :: "i\<Rightarrow>i"  where
     "FP(F) \<equiv> {s\<in>state. F \<in> stable({s})}"
 
 
--- a/src/ZF/UNITY/Follows.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Follows.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -10,7 +10,7 @@
 theory Follows imports SubstAx Increasing begin
 
 definition
-  Follows :: "[i, i, i=>i, i=>i] => i"  where
+  Follows :: "[i, i, i\<Rightarrow>i, i\<Rightarrow>i] \<Rightarrow> i"  where
   "Follows(A, r, f, g) \<equiv>
             Increasing(A, r, g) Int
             Increasing(A, r,f) Int
@@ -18,27 +18,27 @@
            (\<Inter>k \<in> A. {s \<in> state. <k, g(s)>:r} \<longmapsto>w {s \<in> state. <k,f(s)>:r})"
 
 abbreviation
-  Incr :: "[i=>i]=>i" where
+  Incr :: "[i\<Rightarrow>i]\<Rightarrow>i" where
   "Incr(f) \<equiv> Increasing(list(nat), prefix(nat), f)"
 
 abbreviation
-  n_Incr :: "[i=>i]=>i" where
+  n_Incr :: "[i\<Rightarrow>i]\<Rightarrow>i" where
   "n_Incr(f) \<equiv> Increasing(nat, Le, f)"
 
 abbreviation
-  s_Incr :: "[i=>i]=>i" where
+  s_Incr :: "[i\<Rightarrow>i]\<Rightarrow>i" where
   "s_Incr(f) \<equiv> Increasing(Pow(nat), SetLe(nat), f)"
 
 abbreviation
-  m_Incr :: "[i=>i]=>i" where
+  m_Incr :: "[i\<Rightarrow>i]\<Rightarrow>i" where
   "m_Incr(f) \<equiv> Increasing(Mult(nat), MultLe(nat, Le), f)"
 
 abbreviation
-  n_Fols :: "[i=>i, i=>i]=>i"   (infixl \<open>n'_Fols\<close> 65)  where
+  n_Fols :: "[i\<Rightarrow>i, i\<Rightarrow>i]\<Rightarrow>i"   (infixl \<open>n'_Fols\<close> 65)  where
   "f n_Fols g \<equiv> Follows(nat, Le, f, g)"
 
 abbreviation
-  Follows' :: "[i=>i, i=>i, i, i] => i"
+  Follows' :: "[i\<Rightarrow>i, i\<Rightarrow>i, i, i] \<Rightarrow> i"
         (\<open>(_ /Fols _ /Wrt (_ /'/ _))\<close> [60, 0, 0, 60] 60)  where
   "f Fols g Wrt r/A \<equiv> Follows(A,r,f,g)"
 
@@ -114,14 +114,14 @@
 apply (rule single_LeadsTo_I, auto)
 apply (drule_tac x = "g (sa) " and A = B in bspec)
 apply auto
-apply (drule_tac x = "f (sa) " and P = "%j. F \<in> X(j) \<longmapsto>w Y(j)" for X Y in bspec)
+apply (drule_tac x = "f (sa) " and P = "\<lambda>j. F \<in> X(j) \<longmapsto>w Y(j)" for X Y in bspec)
 apply auto
 apply (rule PSP_Stable [THEN LeadsTo_weaken], blast, blast)
 apply auto
 apply (force simp add: part_order_def refl_def)
 apply (force simp add: part_order_def refl_def)
-apply (drule_tac x = "f1 (x)" and x1 = "f (sa) " and P2 = "%x y. \<forall>u\<in>B. P (x,y,u)" for P in bspec [THEN bspec])
-apply (drule_tac [3] x = "g (x) " and x1 = "g (sa) " and P2 = "%x y. P (x,y) \<longrightarrow> d (x,y) \<in> t" for P d in bspec [THEN bspec])
+apply (drule_tac x = "f1 (x)" and x1 = "f (sa) " and P2 = "\<lambda>x y. \<forall>u\<in>B. P (x,y,u)" for P in bspec [THEN bspec])
+apply (drule_tac [3] x = "g (x) " and x1 = "g (sa) " and P2 = "\<lambda>x y. P (x,y) \<longrightarrow> d (x,y) \<in> t" for P d in bspec [THEN bspec])
 apply auto
 apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD)
 apply (auto simp add: part_order_def)
@@ -135,7 +135,7 @@
   F:{x \<in> state. <k, h(f(x), g(x))> \<in> t} \<longmapsto>w {x \<in> state. <k, h(f(x), g1(x))> \<in> t}"
 apply (unfold mono2_def Increasing_def)
 apply (rule single_LeadsTo_I, auto)
-apply (drule_tac x = "f (sa) " and P = "%k. F \<in> Stable (X (k))" for X in bspec)
+apply (drule_tac x = "f (sa) " and P = "\<lambda>k. F \<in> Stable (X (k))" for X in bspec)
 apply auto
 apply (drule_tac x = "g (sa) " in bspec)
 apply auto
@@ -144,7 +144,7 @@
 apply (force simp add: part_order_def refl_def)
 apply (force simp add: part_order_def refl_def)
 apply (drule_tac x = "f (x) " and x1 = "f (sa) " in bspec [THEN bspec])
-apply (drule_tac [3] x = "g1 (x) " and x1 = "g (sa) " and P2 = "%x y. P (x,y) \<longrightarrow> d (x,y) \<in> t" for P d in bspec [THEN bspec])
+apply (drule_tac [3] x = "g1 (x) " and x1 = "g (sa) " and P2 = "\<lambda>x y. P (x,y) \<longrightarrow> d (x,y) \<in> t" for P d in bspec [THEN bspec])
 apply auto
 apply (rule_tac b = "h (f (sa), g (sa))" and A = C in trans_onD)
 apply (auto simp add: part_order_def)
@@ -180,7 +180,7 @@
 done
 
 lemma Follows_constantI:
- "\<lbrakk>F \<in> program; c \<in> A; refl(A, r)\<rbrakk> \<Longrightarrow> F \<in> Follows(A, r, %x. c, %x. c)"
+ "\<lbrakk>F \<in> program; c \<in> A; refl(A, r)\<rbrakk> \<Longrightarrow> F \<in> Follows(A, r, \<lambda>x. c, \<lambda>x. c)"
 apply (unfold Follows_def, auto)
 apply (auto simp add: refl_def)
 done
@@ -210,7 +210,7 @@
 lemma imp_Follows_comp2:
 "\<lbrakk>F \<in> Follows(A, r, f1, f);  F \<in> Follows(B, s, g1, g);
    mono2(A, r, B, s, C, t, h); refl(A,r); refl(B, s); trans[C](t)\<rbrakk>
-   \<Longrightarrow> F \<in> Follows(C, t, %x. h(f1(x), g1(x)), %x. h(f(x), g(x)))"
+   \<Longrightarrow> F \<in> Follows(C, t, \<lambda>x. h(f1(x), g1(x)), \<lambda>x. h(f(x), g(x)))"
 apply (unfold Follows_def, clarify)
 apply (frule_tac f = g in IncreasingD)
 apply (frule_tac f = f in IncreasingD)
@@ -324,13 +324,13 @@
 lemma increasing_Un:
      "\<lbrakk>F \<in> Increasing.increasing(Pow(A), SetLe(A), f);
          F \<in> Increasing.increasing(Pow(A), SetLe(A), g)\<rbrakk>
-     \<Longrightarrow> F \<in> Increasing.increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
+     \<Longrightarrow> F \<in> Increasing.increasing(Pow(A), SetLe(A), \<lambda>x. f(x) \<union> g(x))"
 by (rule_tac h = "(Un)" in imp_increasing_comp2, auto)
 
 lemma Increasing_Un:
      "\<lbrakk>F \<in> Increasing(Pow(A), SetLe(A), f);
          F \<in> Increasing(Pow(A), SetLe(A), g)\<rbrakk>
-     \<Longrightarrow> F \<in> Increasing(Pow(A), SetLe(A), %x. f(x) \<union> g(x))"
+     \<Longrightarrow> F \<in> Increasing(Pow(A), SetLe(A), \<lambda>x. f(x) \<union> g(x))"
 by (rule_tac h = "(Un)" in imp_Increasing_comp2, auto)
 
 lemma Always_Un:
@@ -342,7 +342,7 @@
 lemma Follows_Un:
 "\<lbrakk>F \<in> Follows(Pow(A), SetLe(A), f1, f);
      F \<in> Follows(Pow(A), SetLe(A), g1, g)\<rbrakk>
-     \<Longrightarrow> F \<in> Follows(Pow(A), SetLe(A), %s. f1(s) \<union> g1(s), %s. f(s) \<union> g(s))"
+     \<Longrightarrow> F \<in> Follows(Pow(A), SetLe(A), \<lambda>s. f1(s) \<union> g1(s), \<lambda>s. f(s) \<union> g(s))"
 by (rule_tac h = "(Un)" in imp_Follows_comp2, auto)
 
 (** Multiset union properties (with the MultLe ordering) **)
@@ -351,12 +351,12 @@
 by (unfold MultLe_def refl_def, auto)
 
 lemma MultLe_refl1 [simp]:
- "\<lbrakk>multiset(M); mset_of(M)<=A\<rbrakk> \<Longrightarrow> <M, M> \<in> MultLe(A, r)"
+ "\<lbrakk>multiset(M); mset_of(M)<=A\<rbrakk> \<Longrightarrow> \<langle>M, M\<rangle> \<in> MultLe(A, r)"
 apply (unfold MultLe_def id_def lam_def)
 apply (auto simp add: Mult_iff_multiset)
 done
 
-lemma MultLe_refl2 [simp]: "M \<in> Mult(A) \<Longrightarrow> <M, M> \<in> MultLe(A, r)"
+lemma MultLe_refl2 [simp]: "M \<in> Mult(A) \<Longrightarrow> \<langle>M, M\<rangle> \<in> MultLe(A, r)"
 by (unfold MultLe_def id_def lam_def, auto)
 
 
@@ -371,7 +371,7 @@
 done
 
 lemma MultLe_trans:
-     "\<lbrakk><M,K> \<in> MultLe(A,r); <K,N> \<in> MultLe(A,r)\<rbrakk> \<Longrightarrow> <M,N> \<in> MultLe(A,r)"
+     "\<lbrakk>\<langle>M,K\<rangle> \<in> MultLe(A,r); \<langle>K,N\<rangle> \<in> MultLe(A,r)\<rbrakk> \<Longrightarrow> \<langle>M,N\<rangle> \<in> MultLe(A,r)"
 apply (cut_tac A=A in trans_on_MultLe)
 apply (drule trans_onD, assumption)
 apply (auto dest: MultLe_type [THEN subsetD])
@@ -405,7 +405,7 @@
 done
 
 lemma empty_le_MultLe [simp]:
-"\<lbrakk>multiset(M); mset_of(M)<= A\<rbrakk> \<Longrightarrow> <0, M> \<in> MultLe(A, r)"
+"\<lbrakk>multiset(M); mset_of(M)<= A\<rbrakk> \<Longrightarrow> \<langle>0, M\<rangle> \<in> MultLe(A, r)"
 apply (unfold MultLe_def)
 apply (case_tac "M=0")
 apply (auto simp add: FiniteFun.intros)
@@ -414,11 +414,11 @@
 apply (auto simp add: Mult_iff_multiset)
 done
 
-lemma empty_le_MultLe2 [simp]: "M \<in> Mult(A) \<Longrightarrow> <0, M> \<in> MultLe(A, r)"
+lemma empty_le_MultLe2 [simp]: "M \<in> Mult(A) \<Longrightarrow> \<langle>0, M\<rangle> \<in> MultLe(A, r)"
 by (simp add: Mult_iff_multiset)
 
 lemma munion_mono:
-"\<lbrakk><M, N> \<in> MultLe(A, r); <K, L> \<in> MultLe(A, r)\<rbrakk> \<Longrightarrow>
+"\<lbrakk>\<langle>M, N\<rangle> \<in> MultLe(A, r); \<langle>K, L\<rangle> \<in> MultLe(A, r)\<rbrakk> \<Longrightarrow>
   <M +# K, N +# L> \<in> MultLe(A, r)"
 apply (unfold MultLe_def)
 apply (auto intro: munion_multirel_mono1 munion_multirel_mono2
@@ -428,13 +428,13 @@
 lemma increasing_munion:
      "\<lbrakk>F \<in> Increasing.increasing(Mult(A), MultLe(A,r), f);
          F \<in> Increasing.increasing(Mult(A), MultLe(A,r), g)\<rbrakk>
-     \<Longrightarrow> F \<in> Increasing.increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))"
+     \<Longrightarrow> F \<in> Increasing.increasing(Mult(A),MultLe(A,r), \<lambda>x. f(x) +# g(x))"
 by (rule_tac h = munion in imp_increasing_comp2, auto)
 
 lemma Increasing_munion:
      "\<lbrakk>F \<in> Increasing(Mult(A), MultLe(A,r), f);
          F \<in> Increasing(Mult(A), MultLe(A,r), g)\<rbrakk>
-     \<Longrightarrow> F \<in> Increasing(Mult(A),MultLe(A,r), %x. f(x) +# g(x))"
+     \<Longrightarrow> F \<in> Increasing(Mult(A),MultLe(A,r), \<lambda>x. f(x) +# g(x))"
 by (rule_tac h = munion in imp_Increasing_comp2, auto)
 
 lemma Always_munion:
@@ -449,7 +449,7 @@
 lemma Follows_munion:
 "\<lbrakk>F \<in> Follows(Mult(A), MultLe(A, r), f1, f);
     F \<in> Follows(Mult(A), MultLe(A, r), g1, g)\<rbrakk>
-  \<Longrightarrow> F \<in> Follows(Mult(A), MultLe(A, r), %s. f1(s) +# g1(s), %s. f(s) +# g(s))"
+  \<Longrightarrow> F \<in> Follows(Mult(A), MultLe(A, r), \<lambda>s. f1(s) +# g1(s), \<lambda>s. f(s) +# g(s))"
 by (rule_tac h = munion in imp_Follows_comp2, auto)
 
 (** Used in ClientImp **)
@@ -460,8 +460,8 @@
                         multiset(f(i, s)) \<and> mset_of(f(i, s))<=A ;
    Finite(I); F \<in> program\<rbrakk>
         \<Longrightarrow> F \<in> Follows(Mult(A),
-                        MultLe(A, r), %x. msetsum(%i. f'(i, x), I, A),
-                                      %x. msetsum(%i. f(i,  x), I, A))"
+                        MultLe(A, r), \<lambda>x. msetsum(\<lambda>i. f'(i, x), I, A),
+                                      \<lambda>x. msetsum(\<lambda>i. f(i,  x), I, A))"
 apply (erule rev_mp)
 apply (drule Finite_into_Fin)
 apply (erule Fin_induct)
--- a/src/ZF/UNITY/GenPrefix.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/GenPrefix.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -2,7 +2,7 @@
     Author:     Sidi O Ehmety, Cambridge University Computer Laboratory
     Copyright   2001  University of Cambridge
 
-<xs,ys>:gen_prefix(r)
+\<langle>xs,ys\<rangle>:gen_prefix(r)
   if ys = xs' @ zs where length(xs) = length(xs')
   and corresponding elements of xs, xs' are pairwise related by r
 
@@ -16,11 +16,11 @@
 begin
 
 definition (*really belongs in ZF/Trancl*)
-  part_order :: "[i, i] => o"  where
+  part_order :: "[i, i] \<Rightarrow> o"  where
   "part_order(A, r) \<equiv> refl(A,r) \<and> trans[A](r) \<and> antisym(r)"
 
 consts
-  gen_prefix :: "[i, i] => i"
+  gen_prefix :: "[i, i] \<Rightarrow> i"
 
 inductive
   (* Parameter A is the domain of zs's elements *)
@@ -30,35 +30,35 @@
   intros
     Nil:     "<[],[]>:gen_prefix(A, r)"
 
-    prepend: "\<lbrakk><xs,ys>:gen_prefix(A, r);  <x,y>:r; x \<in> A; y \<in> A\<rbrakk>
+    prepend: "\<lbrakk>\<langle>xs,ys\<rangle>:gen_prefix(A, r);  \<langle>x,y\<rangle>:r; x \<in> A; y \<in> A\<rbrakk>
               \<Longrightarrow> <Cons(x,xs), Cons(y,ys)>: gen_prefix(A, r)"
 
-    append:  "\<lbrakk><xs,ys>:gen_prefix(A, r); zs:list(A)\<rbrakk>
+    append:  "\<lbrakk>\<langle>xs,ys\<rangle>:gen_prefix(A, r); zs:list(A)\<rbrakk>
               \<Longrightarrow> <xs, ys@zs>:gen_prefix(A, r)"
     type_intros app_type list.Nil list.Cons
 
 definition
-  prefix :: "i=>i"  where
+  prefix :: "i\<Rightarrow>i"  where
   "prefix(A) \<equiv> gen_prefix(A, id(A))"
 
 definition
-  strict_prefix :: "i=>i"  where
+  strict_prefix :: "i\<Rightarrow>i"  where
   "strict_prefix(A) \<equiv> prefix(A) - id(list(A))"
 
 
 (* less or equal and greater or equal over prefixes *)
 
 abbreviation
-  pfixLe :: "[i, i] => o"  (infixl \<open>pfixLe\<close> 50)  where
-  "xs pfixLe ys \<equiv> <xs, ys>:gen_prefix(nat, Le)"
+  pfixLe :: "[i, i] \<Rightarrow> o"  (infixl \<open>pfixLe\<close> 50)  where
+  "xs pfixLe ys \<equiv> \<langle>xs, ys\<rangle>:gen_prefix(nat, Le)"
 
 abbreviation
-  pfixGe :: "[i, i] => o"  (infixl \<open>pfixGe\<close> 50)  where
-  "xs pfixGe ys \<equiv> <xs, ys>:gen_prefix(nat, Ge)"
+  pfixGe :: "[i, i] \<Rightarrow> o"  (infixl \<open>pfixGe\<close> 50)  where
+  "xs pfixGe ys \<equiv> \<langle>xs, ys\<rangle>:gen_prefix(nat, Ge)"
 
 
 lemma reflD:
- "\<lbrakk>refl(A, r); x \<in> A\<rbrakk> \<Longrightarrow> <x,x>:r"
+ "\<lbrakk>refl(A, r); x \<in> A\<rbrakk> \<Longrightarrow> \<langle>x,x\<rangle>:r"
 apply (unfold refl_def, auto)
 done
 
@@ -69,7 +69,7 @@
 declare Nil_gen_prefix [simp]
 
 
-lemma gen_prefix_length_le: "<xs,ys> \<in> gen_prefix(A, r) \<Longrightarrow> length(xs) \<le> length(ys)"
+lemma gen_prefix_length_le: "\<langle>xs,ys\<rangle> \<in> gen_prefix(A, r) \<Longrightarrow> length(xs) \<le> length(ys)"
 apply (erule gen_prefix.induct)
 apply (subgoal_tac [3] "ys \<in> list (A) ")
 apply (auto dest: gen_prefix.dom_subset [THEN subsetD]
@@ -81,15 +81,15 @@
   "\<lbrakk><xs', ys'> \<in> gen_prefix(A, r)\<rbrakk>
    \<Longrightarrow> (\<forall>x xs. x \<in> A \<longrightarrow> xs'= Cons(x,xs) \<longrightarrow>
        (\<exists>y ys. y \<in> A \<and> ys' = Cons(y,ys) \<and>
-       <x,y>:r \<and> <xs, ys> \<in> gen_prefix(A, r)))"
+       \<langle>x,y\<rangle>:r \<and> \<langle>xs, ys\<rangle> \<in> gen_prefix(A, r)))"
 apply (erule gen_prefix.induct)
 prefer 3 apply (force intro: gen_prefix.append, auto)
 done
 
 lemma Cons_gen_prefixE:
   "\<lbrakk><Cons(x,xs), zs> \<in> gen_prefix(A, r);
-    \<And>y ys. \<lbrakk>zs = Cons(y, ys); y \<in> A; x \<in> A; <x,y>:r;
-      <xs,ys> \<in> gen_prefix(A, r)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+    \<And>y ys. \<lbrakk>zs = Cons(y, ys); y \<in> A; x \<in> A; \<langle>x,y\<rangle>:r;
+      \<langle>xs,ys\<rangle> \<in> gen_prefix(A, r)\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 apply (frule gen_prefix.dom_subset [THEN subsetD], auto)
 apply (blast dest: Cons_gen_prefix_aux)
 done
@@ -97,7 +97,7 @@
 
 lemma Cons_gen_prefix_Cons:
 "(<Cons(x,xs),Cons(y,ys)> \<in> gen_prefix(A, r))
-  \<longleftrightarrow> (x \<in> A \<and> y \<in> A \<and> <x,y>:r \<and> <xs,ys> \<in> gen_prefix(A, r))"
+  \<longleftrightarrow> (x \<in> A \<and> y \<in> A \<and> \<langle>x,y\<rangle>:r \<and> \<langle>xs,ys\<rangle> \<in> gen_prefix(A, r))"
 apply (auto intro: gen_prefix.prepend)
 done
 declare Cons_gen_prefix_Cons [iff]
@@ -145,7 +145,7 @@
 (* A lemma for proving gen_prefix_trans_comp *)
 
 lemma append_gen_prefix [rule_format (no_asm)]: "xs \<in> list(A) \<Longrightarrow>
-   \<forall>zs. <xs @ ys, zs> \<in> gen_prefix(A, r) \<longrightarrow> <xs, zs>: gen_prefix(A, r)"
+   \<forall>zs. <xs @ ys, zs> \<in> gen_prefix(A, r) \<longrightarrow> \<langle>xs, zs\<rangle>: gen_prefix(A, r)"
 apply (erule list.induct)
 apply (auto dest: gen_prefix.dom_subset [THEN subsetD])
 done
@@ -153,8 +153,8 @@
 (* Lemma proving transitivity and more*)
 
 lemma gen_prefix_trans_comp [rule_format (no_asm)]:
-     "<x, y>: gen_prefix(A, r) \<Longrightarrow>
-   (\<forall>z \<in> list(A). <y,z> \<in> gen_prefix(A, s)\<longrightarrow><x, z> \<in> gen_prefix(A, s O r))"
+     "\<langle>x, y\<rangle>: gen_prefix(A, r) \<Longrightarrow>
+   (\<forall>z \<in> list(A). \<langle>y,z\<rangle> \<in> gen_prefix(A, s)\<longrightarrow>\<langle>x, z\<rangle> \<in> gen_prefix(A, s O r))"
 apply (erule gen_prefix.induct)
 apply (auto elim: ConsE simp add: Nil_gen_prefix)
 apply (subgoal_tac "ys \<in> list (A) ")
@@ -180,19 +180,19 @@
 done
 
 lemma prefix_gen_prefix_trans:
-    "\<lbrakk><x,y> \<in> prefix(A); <y, z> \<in> gen_prefix(A, r); r<=A*A\<rbrakk>
-      \<Longrightarrow>  <x, z> \<in> gen_prefix(A, r)"
+    "\<lbrakk>\<langle>x,y\<rangle> \<in> prefix(A); \<langle>y, z\<rangle> \<in> gen_prefix(A, r); r<=A*A\<rbrakk>
+      \<Longrightarrow>  \<langle>x, z\<rangle> \<in> gen_prefix(A, r)"
 apply (unfold prefix_def)
-apply (rule_tac P = "%r. <x,z> \<in> gen_prefix (A, r) " in right_comp_id [THEN subst])
+apply (rule_tac P = "\<lambda>r. \<langle>x,z\<rangle> \<in> gen_prefix (A, r) " in right_comp_id [THEN subst])
 apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+
 done
 
 
 lemma gen_prefix_prefix_trans:
-"\<lbrakk><x,y> \<in> gen_prefix(A,r); <y, z> \<in> prefix(A); r<=A*A\<rbrakk>
-  \<Longrightarrow>  <x, z> \<in> gen_prefix(A, r)"
+"\<lbrakk>\<langle>x,y\<rangle> \<in> gen_prefix(A,r); \<langle>y, z\<rangle> \<in> prefix(A); r<=A*A\<rbrakk>
+  \<Longrightarrow>  \<langle>x, z\<rangle> \<in> gen_prefix(A, r)"
 apply (unfold prefix_def)
-apply (rule_tac P = "%r. <x,z> \<in> gen_prefix (A, r) " in left_comp_id [THEN subst])
+apply (rule_tac P = "\<lambda>r. \<langle>x,z\<rangle> \<in> gen_prefix (A, r) " in left_comp_id [THEN subst])
 apply (blast dest: gen_prefix_trans_comp gen_prefix.dom_subset [THEN subsetD])+
 done
 
@@ -231,7 +231,7 @@
 
 lemma same_gen_prefix_gen_prefix:
  "\<lbrakk>refl(A, r);  xs \<in> list(A)\<rbrakk> \<Longrightarrow>
-    <xs@ys, xs@zs>: gen_prefix(A, r) \<longleftrightarrow> <ys,zs> \<in> gen_prefix(A, r)"
+    <xs@ys, xs@zs>: gen_prefix(A, r) \<longleftrightarrow> \<langle>ys,zs\<rangle> \<in> gen_prefix(A, r)"
 apply (unfold refl_def)
 apply (induct_tac "xs")
 apply (simp_all (no_asm_simp))
@@ -240,11 +240,11 @@
 
 lemma gen_prefix_Cons: "\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow>
     <xs, Cons(y,ys)> \<in> gen_prefix(A,r)  \<longleftrightarrow>
-      (xs=[] | (\<exists>z zs. xs=Cons(z,zs) \<and> z \<in> A \<and> <z,y>:r \<and> <zs,ys> \<in> gen_prefix(A,r)))"
+      (xs=[] | (\<exists>z zs. xs=Cons(z,zs) \<and> z \<in> A \<and> \<langle>z,y\<rangle>:r \<and> \<langle>zs,ys\<rangle> \<in> gen_prefix(A,r)))"
 apply (induct_tac "xs", auto)
 done
 
-lemma gen_prefix_take_append: "\<lbrakk>refl(A,r);  <xs,ys> \<in> gen_prefix(A, r); zs \<in> list(A)\<rbrakk>
+lemma gen_prefix_take_append: "\<lbrakk>refl(A,r);  \<langle>xs,ys\<rangle> \<in> gen_prefix(A, r); zs \<in> list(A)\<rbrakk>
       \<Longrightarrow>  <xs@zs, take(length(xs), ys) @ zs> \<in> gen_prefix(A, r)"
 apply (erule gen_prefix.induct)
 apply (simp (no_asm_simp))
@@ -254,7 +254,7 @@
 apply (simp_all (no_asm_simp) add: diff_is_0_iff [THEN iffD2] take_type)
 done
 
-lemma gen_prefix_append_both: "\<lbrakk>refl(A, r);  <xs,ys> \<in> gen_prefix(A,r);
+lemma gen_prefix_append_both: "\<lbrakk>refl(A, r);  \<langle>xs,ys\<rangle> \<in> gen_prefix(A,r);
          length(xs) = length(ys); zs \<in> list(A)\<rbrakk>
       \<Longrightarrow>  <xs@zs, ys @ zs> \<in> gen_prefix(A, r)"
 apply (drule_tac zs = zs in gen_prefix_take_append, assumption+)
@@ -267,7 +267,7 @@
 by (auto simp add: app_assoc)
 
 lemma append_one_gen_prefix_lemma [rule_format]:
-     "\<lbrakk><xs,ys> \<in> gen_prefix(A, r);  refl(A, r)\<rbrakk>
+     "\<lbrakk>\<langle>xs,ys\<rangle> \<in> gen_prefix(A, r);  refl(A, r)\<rbrakk>
       \<Longrightarrow> length(xs) < length(ys) \<longrightarrow>
           <xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)"
 apply (erule gen_prefix.induct, blast)
@@ -283,14 +283,14 @@
 apply (blast intro: gen_prefix.append)
 apply (erule_tac V = "length (xs) < length (ys) \<longrightarrow>u" for u in thin_rl)
 apply (erule_tac a = zs in list.cases, auto)
-apply (rule_tac P1 = "%x. <u(x), v>:w" for u v w in nat_diff_split [THEN iffD2])
+apply (rule_tac P1 = "\<lambda>x. <u(x), v>:w" for u v w in nat_diff_split [THEN iffD2])
 apply auto
 apply (simplesubst append_cons_conv)
 apply (rule_tac [2] gen_prefix.append)
 apply (auto elim: ConsE simp add: gen_prefix_append_both)
 done
 
-lemma append_one_gen_prefix: "\<lbrakk><xs,ys>: gen_prefix(A, r);  length(xs) < length(ys);  refl(A, r)\<rbrakk>
+lemma append_one_gen_prefix: "\<lbrakk>\<langle>xs,ys\<rangle>: gen_prefix(A, r);  length(xs) < length(ys);  refl(A, r)\<rbrakk>
       \<Longrightarrow> <xs @ [nth(length(xs), ys)], ys> \<in> gen_prefix(A, r)"
 apply (blast intro: append_one_gen_prefix_lemma)
 done
@@ -300,13 +300,13 @@
 
 lemma gen_prefix_imp_nth_lemma [rule_format]: "xs \<in> list(A) \<Longrightarrow>
   \<forall>ys \<in> list(A). \<forall>i \<in> nat. i < length(xs)
-          \<longrightarrow> <xs, ys>: gen_prefix(A, r) \<longrightarrow> <nth(i, xs), nth(i, ys)>:r"
+          \<longrightarrow> \<langle>xs, ys\<rangle>: gen_prefix(A, r) \<longrightarrow> <nth(i, xs), nth(i, ys)>:r"
 apply (induct_tac "xs", simp, clarify)
 apply simp
 apply (erule natE, auto)
 done
 
-lemma gen_prefix_imp_nth: "\<lbrakk><xs,ys> \<in> gen_prefix(A,r); i < length(xs)\<rbrakk>
+lemma gen_prefix_imp_nth: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> gen_prefix(A,r); i < length(xs)\<rbrakk>
       \<Longrightarrow> <nth(i, xs), nth(i, ys)>:r"
 apply (cut_tac A = A in gen_prefix.dom_subset)
 apply (rule gen_prefix_imp_nth_lemma)
@@ -316,7 +316,7 @@
 lemma nth_imp_gen_prefix [rule_format]: "xs \<in> list(A) \<Longrightarrow>
   \<forall>ys \<in> list(A). length(xs) \<le> length(ys)
       \<longrightarrow> (\<forall>i. i < length(xs) \<longrightarrow> <nth(i, xs), nth(i,ys)>:r)
-      \<longrightarrow> <xs, ys> \<in> gen_prefix(A, r)"
+      \<longrightarrow> \<langle>xs, ys\<rangle> \<in> gen_prefix(A, r)"
 apply (induct_tac "xs")
 apply (simp_all (no_asm_simp))
 apply clarify
@@ -324,7 +324,7 @@
 apply (force intro!: nat_0_le simp add: lt_nat_in_nat)
 done
 
-lemma gen_prefix_iff_nth: "(<xs,ys> \<in> gen_prefix(A,r)) \<longleftrightarrow>
+lemma gen_prefix_iff_nth: "(\<langle>xs,ys\<rangle> \<in> gen_prefix(A,r)) \<longleftrightarrow>
       (xs \<in> list(A) \<and> ys \<in> list(A) \<and> length(xs) \<le> length(ys) \<and>
       (\<forall>i. i < length(xs) \<longrightarrow> <nth(i,xs), nth(i, ys)>: r))"
 apply (rule iffI)
@@ -363,7 +363,7 @@
 (* Monotonicity of "set" operator WRT prefix *)
 
 lemma set_of_list_prefix_mono:
-"<xs,ys> \<in> prefix(A) \<Longrightarrow> set_of_list(xs) \<subseteq> set_of_list(ys)"
+"\<langle>xs,ys\<rangle> \<in> prefix(A) \<Longrightarrow> set_of_list(xs) \<subseteq> set_of_list(ys)"
 
 apply (unfold prefix_def)
 apply (erule gen_prefix.induct)
@@ -392,13 +392,13 @@
 declare prefix_Nil [iff]
 
 lemma Cons_prefix_Cons:
-"<Cons(x,xs), Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow> (x=y \<and> <xs,ys> \<in> prefix(A) \<and> y \<in> A)"
+"<Cons(x,xs), Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow> (x=y \<and> \<langle>xs,ys\<rangle> \<in> prefix(A) \<and> y \<in> A)"
 apply (unfold prefix_def, auto)
 done
 declare Cons_prefix_Cons [iff]
 
 lemma same_prefix_prefix:
-"xs \<in> list(A)\<Longrightarrow> <xs@ys,xs@zs> \<in> prefix(A) \<longleftrightarrow> (<ys,zs> \<in> prefix(A))"
+"xs \<in> list(A)\<Longrightarrow> <xs@ys,xs@zs> \<in> prefix(A) \<longleftrightarrow> (\<langle>ys,zs\<rangle> \<in> prefix(A))"
 apply (unfold prefix_def)
 apply (subgoal_tac "refl (A,id (A))")
 apply (simp (no_asm_simp))
@@ -407,13 +407,13 @@
 declare same_prefix_prefix [simp]
 
 lemma same_prefix_prefix_Nil: "xs \<in> list(A) \<Longrightarrow> <xs@ys,xs> \<in> prefix(A) \<longleftrightarrow> (<ys,[]> \<in> prefix(A))"
-apply (rule_tac P = "%x. <u, x>:v \<longleftrightarrow> w(x)" for u v w in app_right_Nil [THEN subst])
+apply (rule_tac P = "\<lambda>x. \<langle>u, x\<rangle>:v \<longleftrightarrow> w(x)" for u v w in app_right_Nil [THEN subst])
 apply (rule_tac [2] same_prefix_prefix, auto)
 done
 declare same_prefix_prefix_Nil [simp]
 
 lemma prefix_appendI:
-"\<lbrakk><xs,ys> \<in> prefix(A); zs \<in> list(A)\<rbrakk> \<Longrightarrow> <xs,ys@zs> \<in> prefix(A)"
+"\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); zs \<in> list(A)\<rbrakk> \<Longrightarrow> <xs,ys@zs> \<in> prefix(A)"
 apply (unfold prefix_def)
 apply (erule gen_prefix.append, assumption)
 done
@@ -422,13 +422,13 @@
 lemma prefix_Cons:
 "\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow>
   <xs,Cons(y,ys)> \<in> prefix(A) \<longleftrightarrow>
-  (xs=[] | (\<exists>zs. xs=Cons(y,zs) \<and> <zs,ys> \<in> prefix(A)))"
+  (xs=[] | (\<exists>zs. xs=Cons(y,zs) \<and> \<langle>zs,ys\<rangle> \<in> prefix(A)))"
 apply (unfold prefix_def)
 apply (auto simp add: gen_prefix_Cons)
 done
 
 lemma append_one_prefix:
-  "\<lbrakk><xs,ys> \<in> prefix(A); length(xs) < length(ys)\<rbrakk>
+  "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); length(xs) < length(ys)\<rbrakk>
   \<Longrightarrow> <xs @ [nth(length(xs),ys)], ys> \<in> prefix(A)"
 apply (unfold prefix_def)
 apply (subgoal_tac "refl (A, id (A))")
@@ -437,7 +437,7 @@
 done
 
 lemma prefix_length_le:
-"<xs,ys> \<in> prefix(A) \<Longrightarrow> length(xs) \<le> length(ys)"
+"\<langle>xs,ys\<rangle> \<in> prefix(A) \<Longrightarrow> length(xs) \<le> length(ys)"
 apply (unfold prefix_def)
 apply (blast dest: gen_prefix_length_le)
 done
@@ -454,7 +454,7 @@
 done
 
 lemma strict_prefix_length_lt_aux:
-     "<xs,ys> \<in> prefix(A) \<Longrightarrow> xs\<noteq>ys \<longrightarrow> length(xs) < length(ys)"
+     "\<langle>xs,ys\<rangle> \<in> prefix(A) \<Longrightarrow> xs\<noteq>ys \<longrightarrow> length(xs) < length(ys)"
 apply (unfold prefix_def)
 apply (erule gen_prefix.induct, clarify)
 apply (subgoal_tac [!] "ys \<in> list(A) \<and> xs \<in> list(A)")
@@ -467,7 +467,7 @@
 done
 
 lemma strict_prefix_length_lt:
-     "<xs,ys>:strict_prefix(A) \<Longrightarrow> length(xs) < length(ys)"
+     "\<langle>xs,ys\<rangle>:strict_prefix(A) \<Longrightarrow> length(xs) < length(ys)"
 apply (unfold strict_prefix_def)
 apply (rule strict_prefix_length_lt_aux [THEN mp])
 apply (auto dest: prefix_type [THEN subsetD])
@@ -475,7 +475,7 @@
 
 (*Equivalence to the definition used in Lex/Prefix.thy*)
 lemma prefix_iff:
-    "<xs,zs> \<in> prefix(A) \<longleftrightarrow> (\<exists>ys \<in> list(A). zs = xs@ys) \<and> xs \<in> list(A)"
+    "\<langle>xs,zs\<rangle> \<in> prefix(A) \<longleftrightarrow> (\<exists>ys \<in> list(A). zs = xs@ys) \<and> xs \<in> list(A)"
 apply (unfold prefix_def)
 apply (auto simp add: gen_prefix_iff_nth lt_nat_in_nat nth_append nth_type app_type length_app)
 apply (subgoal_tac "drop (length (xs), zs) \<in> list (A) ")
@@ -494,7 +494,7 @@
 
 lemma prefix_snoc:
 "\<lbrakk>xs \<in> list(A); ys \<in> list(A); y \<in> A\<rbrakk> \<Longrightarrow>
-   <xs, ys@[y]> \<in> prefix(A) \<longleftrightarrow> (xs = ys@[y] | <xs,ys> \<in> prefix(A))"
+   <xs, ys@[y]> \<in> prefix(A) \<longleftrightarrow> (xs = ys@[y] | \<langle>xs,ys\<rangle> \<in> prefix(A))"
 apply (simp (no_asm) add: prefix_iff)
 apply (rule iffI, clarify)
 apply (erule_tac xs = ysa in rev_list_elim, simp)
@@ -505,7 +505,7 @@
 
 lemma prefix_append_iff [rule_format]: "zs \<in> list(A) \<Longrightarrow> \<forall>xs \<in> list(A). \<forall>ys \<in> list(A).
    (<xs, ys@zs> \<in> prefix(A)) \<longleftrightarrow>
-  (<xs,ys> \<in> prefix(A) | (\<exists>us. xs = ys@us \<and> <us,zs> \<in> prefix(A)))"
+  (\<langle>xs,ys\<rangle> \<in> prefix(A) | (\<exists>us. xs = ys@us \<and> \<langle>us,zs\<rangle> \<in> prefix(A)))"
 apply (erule list_append_induct, force, clarify)
 apply (rule iffI)
 apply (simp add: add: app_assoc [symmetric])
@@ -519,13 +519,13 @@
 (*Although the prefix ordering is not linear, the prefixes of a list
   are linearly ordered.*)
 lemma common_prefix_linear_lemma [rule_format]: "\<lbrakk>zs \<in> list(A); xs \<in> list(A); ys \<in> list(A)\<rbrakk>
-   \<Longrightarrow> <xs, zs> \<in> prefix(A) \<longrightarrow> <ys,zs> \<in> prefix(A)
-  \<longrightarrow><xs,ys> \<in> prefix(A) | <ys,xs> \<in> prefix(A)"
+   \<Longrightarrow> \<langle>xs, zs\<rangle> \<in> prefix(A) \<longrightarrow> \<langle>ys,zs\<rangle> \<in> prefix(A)
+  \<longrightarrow>\<langle>xs,ys\<rangle> \<in> prefix(A) | \<langle>ys,xs\<rangle> \<in> prefix(A)"
 apply (erule list_append_induct, auto)
 done
 
-lemma common_prefix_linear: "\<lbrakk><xs, zs> \<in> prefix(A); <ys,zs> \<in> prefix(A)\<rbrakk>
-      \<Longrightarrow> <xs,ys> \<in> prefix(A) | <ys,xs> \<in> prefix(A)"
+lemma common_prefix_linear: "\<lbrakk>\<langle>xs, zs\<rangle> \<in> prefix(A); \<langle>ys,zs\<rangle> \<in> prefix(A)\<rbrakk>
+      \<Longrightarrow> \<langle>xs,ys\<rangle> \<in> prefix(A) | \<langle>ys,xs\<rangle> \<in> prefix(A)"
 apply (cut_tac prefix_type)
 apply (blast del: disjCI intro: common_prefix_linear_lemma)
 done
@@ -577,7 +577,7 @@
 
 
 lemma prefix_imp_pfixLe:
-"<xs,ys>:prefix(nat)\<Longrightarrow> xs pfixLe ys"
+"\<langle>xs,ys\<rangle>:prefix(nat)\<Longrightarrow> xs pfixLe ys"
 
 apply (unfold prefix_def)
 apply (rule gen_prefix_mono [THEN subsetD], auto)
@@ -610,14 +610,14 @@
 by (blast intro: antisym_gen_prefix [THEN antisymE])
 
 lemma prefix_imp_pfixGe:
-  "<xs,ys>:prefix(nat) \<Longrightarrow> xs pfixGe ys"
+  "\<langle>xs,ys\<rangle>:prefix(nat) \<Longrightarrow> xs pfixGe ys"
 apply (unfold prefix_def Ge_def)
 apply (rule gen_prefix_mono [THEN subsetD], auto)
 done
 (* Added by Sidi \<in> prefix and take *)
 
 lemma prefix_imp_take:
-"<xs, ys> \<in> prefix(A) \<Longrightarrow> xs = take(length(xs), ys)"
+"\<langle>xs, ys\<rangle> \<in> prefix(A) \<Longrightarrow> xs = take(length(xs), ys)"
 
 apply (unfold prefix_def)
 apply (erule gen_prefix.induct)
@@ -630,14 +630,14 @@
 apply (simp_all (no_asm_simp) add: diff_is_0_iff)
 done
 
-lemma prefix_length_equal: "\<lbrakk><xs,ys> \<in> prefix(A); length(xs)=length(ys)\<rbrakk> \<Longrightarrow> xs = ys"
+lemma prefix_length_equal: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); length(xs)=length(ys)\<rbrakk> \<Longrightarrow> xs = ys"
 apply (cut_tac A = A in prefix_type)
 apply (drule subsetD, auto)
 apply (drule prefix_imp_take)
 apply (erule trans, simp)
 done
 
-lemma prefix_length_le_equal: "\<lbrakk><xs,ys> \<in> prefix(A); length(ys) \<le> length(xs)\<rbrakk> \<Longrightarrow> xs = ys"
+lemma prefix_length_le_equal: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); length(ys) \<le> length(xs)\<rbrakk> \<Longrightarrow> xs = ys"
 by (blast intro: prefix_length_equal le_anti_sym prefix_length_le)
 
 lemma take_prefix [rule_format]: "xs \<in> list(A) \<Longrightarrow> \<forall>n \<in> nat. <take(n, xs), xs> \<in> prefix(A)"
@@ -646,7 +646,7 @@
 apply (erule natE, auto)
 done
 
-lemma prefix_take_iff: "<xs,ys> \<in> prefix(A) \<longleftrightarrow> (xs=take(length(xs), ys) \<and> xs \<in> list(A) \<and> ys \<in> list(A))"
+lemma prefix_take_iff: "\<langle>xs,ys\<rangle> \<in> prefix(A) \<longleftrightarrow> (xs=take(length(xs), ys) \<and> xs \<in> list(A) \<and> ys \<in> list(A))"
 apply (rule iffI)
 apply (frule prefix_type [THEN subsetD])
 apply (blast intro: prefix_imp_take, clarify)
@@ -654,13 +654,13 @@
 apply (blast intro: take_prefix length_type)
 done
 
-lemma prefix_imp_nth: "\<lbrakk><xs,ys> \<in> prefix(A); i < length(xs)\<rbrakk> \<Longrightarrow> nth(i,xs) = nth(i,ys)"
+lemma prefix_imp_nth: "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); i < length(xs)\<rbrakk> \<Longrightarrow> nth(i,xs) = nth(i,ys)"
 by (auto dest!: gen_prefix_imp_nth simp add: prefix_def)
 
 lemma nth_imp_prefix:
      "\<lbrakk>xs \<in> list(A); ys \<in> list(A); length(xs) \<le> length(ys);
         \<And>i. i < length(xs) \<Longrightarrow> nth(i, xs) = nth(i,ys)\<rbrakk>
-      \<Longrightarrow> <xs,ys> \<in> prefix(A)"
+      \<Longrightarrow> \<langle>xs,ys\<rangle> \<in> prefix(A)"
 apply (auto simp add: prefix_def nth_imp_gen_prefix)
 apply (auto intro!: nth_imp_gen_prefix simp add: prefix_def)
 apply (blast intro: nth_type lt_trans2)
@@ -668,7 +668,7 @@
 
 
 lemma length_le_prefix_imp_prefix: "\<lbrakk>length(xs) \<le> length(ys);
-        <xs,zs> \<in> prefix(A); <ys,zs> \<in> prefix(A)\<rbrakk> \<Longrightarrow> <xs,ys> \<in> prefix(A)"
+        \<langle>xs,zs\<rangle> \<in> prefix(A); \<langle>ys,zs\<rangle> \<in> prefix(A)\<rbrakk> \<Longrightarrow> \<langle>xs,ys\<rangle> \<in> prefix(A)"
 apply (cut_tac A = A in prefix_type)
 apply (rule nth_imp_prefix, blast, blast)
  apply assumption
--- a/src/ZF/UNITY/Guar.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Guar.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -37,38 +37,38 @@
     case, proving equivalence with Chandy and Sanders's n-ary definitions*)
 
 definition
-   ex_prop :: "i => o"  where
+   ex_prop :: "i \<Rightarrow> o"  where
    "ex_prop(X) \<equiv> X<=program \<and>
     (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X | G \<in> X \<longrightarrow> (F \<squnion> G) \<in> X)"
 
 definition
-  strict_ex_prop  :: "i => o"  where
+  strict_ex_prop  :: "i \<Rightarrow> o"  where
   "strict_ex_prop(X) \<equiv> X<=program \<and>
    (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> (F \<in> X | G \<in> X) \<longleftrightarrow> (F \<squnion> G \<in> X))"
 
 definition
-  uv_prop  :: "i => o"  where
+  uv_prop  :: "i \<Rightarrow> o"  where
    "uv_prop(X) \<equiv> X<=program \<and>
     (SKIP \<in> X \<and> (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<in> X \<and> G \<in> X \<longrightarrow> (F \<squnion> G) \<in> X))"
   
 definition
- strict_uv_prop  :: "i => o"  where
+ strict_uv_prop  :: "i \<Rightarrow> o"  where
    "strict_uv_prop(X) \<equiv> X<=program \<and>
     (SKIP \<in> X \<and> (\<forall>F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow>(F \<in> X \<and> G \<in> X) \<longleftrightarrow> (F \<squnion> G \<in> X)))"
 
 definition
-  guar :: "[i, i] => i" (infixl \<open>guarantees\<close> 55)  where
+  guar :: "[i, i] \<Rightarrow> i" (infixl \<open>guarantees\<close> 55)  where
               (*higher than membership, lower than Co*)
   "X guarantees Y \<equiv> {F \<in> program. \<forall>G \<in> program. F ok G \<longrightarrow> F \<squnion> G \<in> X \<longrightarrow> F \<squnion> G \<in> Y}"
   
 definition
   (* Weakest guarantees *)
-   wg :: "[i,i] => i"  where
+   wg :: "[i,i] \<Rightarrow> i"  where
   "wg(F,Y) \<equiv> \<Union>({X \<in> Pow(program). F:(X guarantees Y)})"
 
 definition
   (* Weakest existential property stronger than X *)
-   wx :: "i =>i"  where
+   wx :: "i \<Rightarrow>i"  where
    "wx(X) \<equiv> \<Union>({Y \<in> Pow(program). Y<=X \<and> ex_prop(Y)})"
 
 definition
@@ -77,13 +77,13 @@
   "welldef \<equiv> {F \<in> program. Init(F) \<noteq> 0}"
   
 definition
-  refines :: "[i, i, i] => o" (\<open>(3_ refines _ wrt _)\<close> [10,10,10] 10)  where
+  refines :: "[i, i, i] \<Rightarrow> o" (\<open>(3_ refines _ wrt _)\<close> [10,10,10] 10)  where
   "G refines F wrt X \<equiv>
    \<forall>H \<in> program. (F ok H  \<and> G ok H \<and> F \<squnion> H \<in> welldef \<inter> X)
     \<longrightarrow> (G \<squnion> H \<in> welldef \<inter> X)"
 
 definition
-  iso_refines :: "[i,i, i] => o"  (\<open>(3_ iso'_refines _ wrt _)\<close> [10,10,10] 10)  where
+  iso_refines :: "[i,i, i] \<Rightarrow> o"  (\<open>(3_ iso'_refines _ wrt _)\<close> [10,10,10] 10)  where
   "G iso_refines F wrt X \<equiv>  F \<in> welldef \<inter> X \<longrightarrow> G \<in> welldef \<inter> X"
 
 
@@ -94,7 +94,7 @@
 
 lemma ex1 [rule_format]: 
  "GG \<in> Fin(program) 
-  \<Longrightarrow> ex_prop(X) \<longrightarrow> GG \<inter> X\<noteq>0 \<longrightarrow> OK(GG, (%G. G)) \<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X"
+  \<Longrightarrow> ex_prop(X) \<longrightarrow> GG \<inter> X\<noteq>0 \<longrightarrow> OK(GG, (\<lambda>G. G)) \<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X"
 apply (unfold ex_prop_def)
 apply (erule Fin_induct)
 apply (simp_all add: OK_cons_iff)
@@ -103,7 +103,7 @@
 
 lemma ex2 [rule_format]: 
 "X \<subseteq> program \<Longrightarrow>  
- (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<longrightarrow> OK(GG,(%G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X) 
+ (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<longrightarrow> OK(GG,(\<lambda>G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X) 
    \<longrightarrow> ex_prop(X)"
 apply (unfold ex_prop_def, clarify)
 apply (drule_tac x = "{F,G}" in bspec)
@@ -115,7 +115,7 @@
 
 lemma ex_prop_finite:
      "ex_prop(X) \<longleftrightarrow> (X\<subseteq>program \<and>  
-  (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<and> OK(GG,(%G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X))"
+  (\<forall>GG \<in> Fin(program). GG \<inter> X \<noteq> 0 \<and> OK(GG,(\<lambda>G. G))\<longrightarrow>(\<Squnion>G \<in> GG. G) \<in> X))"
 apply auto
 apply (blast intro: ex1 ex2 dest: ex_imp_subset_program)+
 done
@@ -138,7 +138,7 @@
 
 lemma uv1 [rule_format]: 
      "GG \<in> Fin(program) \<Longrightarrow>  
-      (uv_prop(X)\<longrightarrow> GG \<subseteq> X \<and> OK(GG, (%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)"
+      (uv_prop(X)\<longrightarrow> GG \<subseteq> X \<and> OK(GG, (\<lambda>G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)"
 apply (unfold uv_prop_def)
 apply (erule Fin_induct)
 apply (auto simp add: OK_cons_iff)
@@ -146,7 +146,7 @@
 
 lemma uv2 [rule_format]: 
      "X\<subseteq>program  \<Longrightarrow> 
-      (\<forall>GG \<in> Fin(program). GG \<subseteq> X \<and> OK(GG,(%G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)
+      (\<forall>GG \<in> Fin(program). GG \<subseteq> X \<and> OK(GG,(\<lambda>G. G)) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in> X)
       \<longrightarrow> uv_prop(X)"
 apply (unfold uv_prop_def, auto)
 apply (drule_tac x = 0 in bspec, simp+) 
@@ -157,7 +157,7 @@
 (*Chandy \<and> Sanders take this as a definition*)
 lemma uv_prop_finite: 
 "uv_prop(X) \<longleftrightarrow> X\<subseteq>program \<and>  
-  (\<forall>GG \<in> Fin(program). GG \<subseteq> X \<and> OK(GG, %G. G) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in>  X)"
+  (\<forall>GG \<in> Fin(program). GG \<subseteq> X \<and> OK(GG, \<lambda>G. G) \<longrightarrow> (\<Squnion>G \<in> GG. G) \<in>  X)"
 apply auto
 apply (blast dest: uv_imp_subset_program)
 apply (blast intro: uv1)
@@ -427,7 +427,7 @@
 by (simp (no_asm_simp) add: wg_equiv)
 
 lemma wg_finite [rule_format]: 
-"\<forall>FF \<in> Fin(program). FF \<inter> X \<noteq> 0 \<longrightarrow> OK(FF, %F. F)  
+"\<forall>FF \<in> Fin(program). FF \<inter> X \<noteq> 0 \<longrightarrow> OK(FF, \<lambda>F. F)  
    \<longrightarrow> (\<forall>F \<in> FF. ((\<Squnion>F \<in> FF. F) \<in>  wg(F,X)) \<longleftrightarrow> ((\<Squnion>F \<in> FF. F) \<in> X))"
 apply clarify
 apply (subgoal_tac "F component_of (\<Squnion>F \<in> FF. F) ")
--- a/src/ZF/UNITY/Increasing.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Increasing.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -11,19 +11,19 @@
 theory Increasing imports Constrains Monotonicity begin
 
 definition
-  increasing :: "[i, i, i=>i] => i" (\<open>increasing[_]'(_, _')\<close>)  where
+  increasing :: "[i, i, i\<Rightarrow>i] \<Rightarrow> i" (\<open>increasing[_]'(_, _')\<close>)  where
   "increasing[A](r, f) \<equiv>
     {F \<in> program. (\<forall>k \<in> A. F \<in> stable({s \<in> state. <k, f(s)> \<in> r})) \<and>
                 (\<forall>x \<in> state. f(x):A)}"
   
 definition
-  Increasing :: "[i, i, i=>i] => i" (\<open>Increasing[_]'(_, _')\<close>)  where
+  Increasing :: "[i, i, i\<Rightarrow>i] \<Rightarrow> i" (\<open>Increasing[_]'(_, _')\<close>)  where
   "Increasing[A](r, f) \<equiv>
     {F \<in> program. (\<forall>k \<in> A. F \<in> Stable({s \<in> state. <k, f(s)> \<in> r})) \<and>
                 (\<forall>x \<in> state. f(x):A)}"
 
 abbreviation (input)
-  IncWrt ::  "[i=>i, i, i] => i" (\<open>(_ IncreasingWrt _ '/ _)\<close> [60, 0, 60] 60)  where
+  IncWrt ::  "[i\<Rightarrow>i, i, i] \<Rightarrow> i" (\<open>(_ IncreasingWrt _ '/ _)\<close> [60, 0, 60] 60)  where
   "f IncreasingWrt r/A \<equiv> Increasing[A](r,f)"
 
 
@@ -47,7 +47,7 @@
 done
 
 lemma increasing_constant [simp]: 
- "F \<in> increasing[A](r, %s. c) \<longleftrightarrow> F \<in> program \<and> c \<in> A"
+ "F \<in> increasing[A](r, \<lambda>s. c) \<longleftrightarrow> F \<in> program \<and> c \<in> A"
 apply (unfold increasing_def stable_def)
 apply (subgoal_tac "\<exists>x. x \<in> state")
 apply (auto dest: stable_type [THEN subsetD] intro: st0_in_state)
@@ -116,7 +116,7 @@
 done
 
 lemma Increasing_constant [simp]: 
-     "F \<in> Increasing[A](r, %s. c) \<longleftrightarrow> F \<in> program \<and> (c \<in> A)"
+     "F \<in> Increasing[A](r, \<lambda>s. c) \<longleftrightarrow> F \<in> program \<and> (c \<in> A)"
 apply (subgoal_tac "\<exists>x. x \<in> state")
 apply (auto dest!: IncreasingD intro: st0_in_state increasing_imp_Increasing)
 done
@@ -163,7 +163,7 @@
 lemma imp_increasing_comp2: 
 "\<lbrakk>F \<in> increasing[A](r, f); F \<in> increasing[B](s, g);  
     mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\<rbrakk>
- \<Longrightarrow> F \<in> increasing[C](t, %x. h(f(x), g(x)))"
+ \<Longrightarrow> F \<in> increasing[C](t, \<lambda>x. h(f(x), g(x)))"
 apply (unfold increasing_def stable_def 
        part_order_def constrains_def mono2_def, clarify, simp)
 apply clarify
@@ -195,7 +195,7 @@
 lemma imp_Increasing_comp2: 
 "\<lbrakk>F \<in> Increasing[A](r, f); F \<in> Increasing[B](s, g);  
   mono2(A, r, B, s, C, t, h); refl(A, r); refl(B, s); trans[C](t)\<rbrakk> \<Longrightarrow>  
-  F \<in> Increasing[C](t, %x. h(f(x), g(x)))"
+  F \<in> Increasing[C](t, \<lambda>x. h(f(x), g(x)))"
 apply (unfold Increasing_def stable_def 
        part_order_def constrains_def mono2_def Stable_def Constrains_def, safe)
 apply (simp_all add: ActsD)
--- a/src/ZF/UNITY/Merge.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Merge.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -13,50 +13,50 @@
 
 definition
   (*spec (10)*)
-  merge_increasing :: "[i, i, i] =>i"  where
+  merge_increasing :: "[i, i, i] \<Rightarrow>i"  where
     "merge_increasing(A, Out, iOut) \<equiv> program guarantees
          (lift(Out) IncreasingWrt  prefix(A)/list(A)) Int
          (lift(iOut) IncreasingWrt prefix(nat)/list(nat))"
 
 definition
   (*spec (11)*)
-  merge_eq_Out :: "[i, i] =>i"  where
+  merge_eq_Out :: "[i, i] \<Rightarrow>i"  where
   "merge_eq_Out(Out, iOut) \<equiv> program guarantees
          Always({s \<in> state. length(s`Out) = length(s`iOut)})"
 
 definition
   (*spec (12)*)
-  merge_bounded :: "i=>i"  where
+  merge_bounded :: "i\<Rightarrow>i"  where
   "merge_bounded(iOut) \<equiv> program guarantees
          Always({s \<in> state. \<forall>elt \<in> set_of_list(s`iOut). elt<Nclients})"
   
 definition
   (*spec (13)*)
   (* Parameter A represents the type of tokens *)
-  merge_follows :: "[i, i=>i, i, i] =>i"  where
+  merge_follows :: "[i, i\<Rightarrow>i, i, i] \<Rightarrow>i"  where
     "merge_follows(A, In, Out, iOut) \<equiv>
      (\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A))
                    guarantees
      (\<Inter>n \<in> Nclients. 
-        (%s. sublist(s`Out, {k \<in> nat. k < length(s`iOut) \<and>
+        (\<lambda>s. sublist(s`Out, {k \<in> nat. k < length(s`iOut) \<and>
                       nth(k, s`iOut) = n})) Fols lift(In(n))
          Wrt prefix(A)/list(A))"
 
 definition
   (*spec: preserves part*)
-  merge_preserves :: "[i=>i] =>i"  where
+  merge_preserves :: "[i\<Rightarrow>i] \<Rightarrow>i"  where
     "merge_preserves(In) \<equiv> \<Inter>n \<in> nat. preserves(lift(In(n)))"
 
 definition
 (* environmental constraints*)
-  merge_allowed_acts :: "[i, i] =>i"  where
+  merge_allowed_acts :: "[i, i] \<Rightarrow>i"  where
   "merge_allowed_acts(Out, iOut) \<equiv>
          {F \<in> program. AllowedActs(F) =
             cons(id(state), (\<Union>G \<in> preserves(lift(Out)) \<inter>
                                    preserves(lift(iOut)). Acts(G)))}"
   
 definition
-  merge_spec :: "[i, i =>i, i, i]=>i"  where
+  merge_spec :: "[i, i \<Rightarrow>i, i, i]\<Rightarrow>i"  where
   "merge_spec(A, In, Out, iOut) \<equiv>
    merge_increasing(A, Out, iOut) \<inter> merge_eq_Out(Out, iOut) \<inter>
    merge_bounded(iOut) \<inter>  merge_follows(A, In, Out, iOut)
@@ -146,7 +146,7 @@
 "\<lbrakk>G \<in> preserves(lift(iOut));  
     G: preserves(lift(Out)); M \<in> Allowed(G)\<rbrakk>  
   \<Longrightarrow> M \<squnion> G \<in> Always  
-    ({s \<in> state. msetsum(%i. bag_of(sublist(s`Out,  
+    ({s \<in> state. msetsum(\<lambda>i. bag_of(sublist(s`Out,  
       {k \<in> nat. k < length(s`iOut) \<and> nth(k, s`iOut)=i})),  
                    Nclients, A) = bag_of(s`Out)})"
 apply (rule Always_Diff_Un_eq [THEN iffD1]) 
@@ -166,7 +166,7 @@
 apply (subgoal_tac "xa \<in> nat")
 apply (simp_all add: Ord_mem_iff_lt)
 prefer 2 apply (blast intro: lt_trans)
-apply (drule_tac x = "nth (xa, x`iOut)" and P = "%elt. X (elt) \<longrightarrow> elt<Nclients" for X in bspec)
+apply (drule_tac x = "nth (xa, x`iOut)" and P = "\<lambda>elt. X (elt) \<longrightarrow> elt<Nclients" for X in bspec)
 apply (simp add: ltI nat_into_Ord)
 apply (blast dest: ltD)
 done
@@ -174,8 +174,8 @@
 theorem (in merge) merge_bag_Follows: 
     "M \<in> (\<Inter>n \<in> Nclients. lift(In(n)) IncreasingWrt prefix(A)/list(A))  
             guarantees   
-      (%s. bag_of(s`Out)) Fols  
-      (%s. msetsum(%i. bag_of(s`In(i)),Nclients, A)) Wrt MultLe(A, r)/Mult(A)"
+      (\<lambda>s. bag_of(s`Out)) Fols  
+      (\<lambda>s. msetsum(\<lambda>i. bag_of(s`In(i)),Nclients, A)) Wrt MultLe(A, r)/Mult(A)"
 apply (cut_tac merge_spec)
 apply (rule merge_bag_Follows_lemma [THEN Always_Follows1, THEN guaranteesI])
      apply (simp_all add: M_ok_iff, clarify)
--- a/src/ZF/UNITY/Monotonicity.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Monotonicity.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,37 +12,37 @@
 begin
 
 definition
-  mono1 :: "[i, i, i, i, i=>i] => o"  where
+  mono1 :: "[i, i, i, i, i\<Rightarrow>i] \<Rightarrow> o"  where
   "mono1(A, r, B, s, f) \<equiv>
-    (\<forall>x \<in> A. \<forall>y \<in> A. <x,y> \<in> r \<longrightarrow> <f(x), f(y)> \<in> s) \<and> (\<forall>x \<in> A. f(x) \<in> B)"
+    (\<forall>x \<in> A. \<forall>y \<in> A. \<langle>x,y\<rangle> \<in> r \<longrightarrow> <f(x), f(y)> \<in> s) \<and> (\<forall>x \<in> A. f(x) \<in> B)"
 
   (* monotonicity of a 2-place meta-function f *)
 
 definition
-  mono2 :: "[i, i, i, i, i, i, [i,i]=>i] => o"  where
+  mono2 :: "[i, i, i, i, i, i, [i,i]\<Rightarrow>i] \<Rightarrow> o"  where
   "mono2(A, r, B, s, C, t, f) \<equiv> 
     (\<forall>x \<in> A. \<forall>y \<in> A. \<forall>u \<in> B. \<forall>v \<in> B.
-              <x,y> \<in> r \<and> <u,v> \<in> s \<longrightarrow> <f(x,u), f(y,v)> \<in> t) \<and>
+              \<langle>x,y\<rangle> \<in> r \<and> \<langle>u,v\<rangle> \<in> s \<longrightarrow> <f(x,u), f(y,v)> \<in> t) \<and>
     (\<forall>x \<in> A. \<forall>y \<in> B. f(x,y) \<in> C)"
 
  (* Internalized relations on sets and multisets *)
 
 definition
-  SetLe :: "i =>i"  where
-  "SetLe(A) \<equiv> {<x,y> \<in> Pow(A)*Pow(A). x \<subseteq> y}"
+  SetLe :: "i \<Rightarrow>i"  where
+  "SetLe(A) \<equiv> {\<langle>x,y\<rangle> \<in> Pow(A)*Pow(A). x \<subseteq> y}"
 
 definition
-  MultLe :: "[i,i] =>i"  where
+  MultLe :: "[i,i] \<Rightarrow>i"  where
   "MultLe(A, r) \<equiv> multirel(A, r - id(A)) \<union> id(Mult(A))"
 
 
 lemma mono1D: 
-  "\<lbrakk>mono1(A, r, B, s, f); <x, y> \<in> r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> <f(x), f(y)> \<in> s"
+  "\<lbrakk>mono1(A, r, B, s, f); \<langle>x, y\<rangle> \<in> r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> <f(x), f(y)> \<in> s"
 by (unfold mono1_def, auto)
 
 lemma mono2D: 
      "\<lbrakk>mono2(A, r, B, s, C, t, f);  
-         <x, y> \<in> r; <u,v> \<in> s; x \<in> A; y \<in> A; u \<in> B; v \<in> B\<rbrakk> 
+         \<langle>x, y\<rangle> \<in> r; \<langle>u,v\<rangle> \<in> s; x \<in> A; y \<in> A; u \<in> B; v \<in> B\<rbrakk> 
       \<Longrightarrow> <f(x, u), f(y,v)> \<in> t"
 by (unfold mono2_def, auto)
 
@@ -72,12 +72,12 @@
 by (blast intro: le_in_nat take_mono_left_lemma) 
 
 lemma take_mono_right:
-     "\<lbrakk><xs,ys> \<in> prefix(A); i \<in> nat\<rbrakk> 
+     "\<lbrakk>\<langle>xs,ys\<rangle> \<in> prefix(A); i \<in> nat\<rbrakk> 
       \<Longrightarrow> <take(i, xs), take(i, ys)> \<in> prefix(A)"
 by (auto simp add: prefix_iff)
 
 lemma take_mono:
-     "\<lbrakk>i \<le> j; <xs, ys> \<in> prefix(A); j \<in> nat\<rbrakk>
+     "\<lbrakk>i \<le> j; \<langle>xs, ys\<rangle> \<in> prefix(A); j \<in> nat\<rbrakk>
       \<Longrightarrow> <take(i, xs), take(j, ys)> \<in> prefix(A)"
 apply (rule_tac b = "take (j, xs) " in prefix_trans)
 apply (auto dest: prefix_type [THEN subsetD] intro: take_mono_left take_mono_right)
--- a/src/ZF/UNITY/MultisetSum.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/MultisetSum.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -9,23 +9,23 @@
 begin
 
 definition
-  lcomm :: "[i, i, [i,i]=>i]=>o"  where
+  lcomm :: "[i, i, [i,i]\<Rightarrow>i]\<Rightarrow>o"  where
   "lcomm(A, B, f) \<equiv>
    (\<forall>x \<in> A. \<forall>y \<in> A. \<forall>z \<in> B. f(x, f(y, z))= f(y, f(x, z)))  \<and>
    (\<forall>x \<in> A. \<forall>y \<in> B. f(x, y) \<in> B)"
 
 definition
-  general_setsum :: "[i,i,i, [i,i]=>i, i=>i] =>i"  where
+  general_setsum :: "[i,i,i, [i,i]\<Rightarrow>i, i\<Rightarrow>i] \<Rightarrow>i"  where
    "general_setsum(C, B, e, f, g) \<equiv>
-    if Finite(C) then fold[cons(e, B)](%x y. f(g(x), y), e, C) else e"
+    if Finite(C) then fold[cons(e, B)](\<lambda>x y. f(g(x), y), e, C) else e"
 
 definition
-  msetsum :: "[i=>i, i, i]=>i"  where
+  msetsum :: "[i\<Rightarrow>i, i, i]\<Rightarrow>i"  where
   "msetsum(g, C, B) \<equiv> normalize(general_setsum(C, Mult(B), 0, (+#), g))"
 
 
 definition  (* sum for naturals *)
-  nsetsum :: "[i=>i, i] =>i"  where
+  nsetsum :: "[i\<Rightarrow>i, i] \<Rightarrow>i"  where
   "nsetsum(g, C) \<equiv> general_setsum(C, nat, 0, (#+), g)"
 
 
@@ -121,7 +121,7 @@
      "\<lbrakk>I \<in> Fin(K); \<forall>i \<in> K. C(i) \<in> Fin(A)\<rbrakk> \<Longrightarrow>  
       (\<forall>x \<in> A. multiset(f(x)) \<and> mset_of(f(x))\<subseteq>B) \<longrightarrow>   
       (\<forall>i \<in> I. \<forall>j \<in> I. i\<noteq>j \<longrightarrow> C(i) \<inter> C(j) = 0) \<longrightarrow>  
-        msetsum(f, \<Union>i \<in> I. C(i), B) = msetsum (%i. msetsum(f, C(i),B), I, B)"
+        msetsum(f, \<Union>i \<in> I. C(i), B) = msetsum (\<lambda>i. msetsum(f, C(i),B), I, B)"
 apply (erule Fin_induct, auto)
 apply (subgoal_tac "\<forall>i \<in> y. x \<noteq> i")
  prefer 2 apply blast
@@ -141,7 +141,7 @@
     "\<lbrakk>C \<in> Fin(A);  
       \<forall>x \<in> A. multiset(f(x)) \<and> mset_of(f(x))\<subseteq>B;  
       \<forall>x \<in> A. multiset(g(x)) \<and> mset_of(g(x))\<subseteq>B\<rbrakk> \<Longrightarrow> 
-      msetsum(%x. f(x) +# g(x), C, B) = msetsum(f, C, B) +# msetsum(g, C, B)"
+      msetsum(\<lambda>x. f(x) +# g(x), C, B) = msetsum(f, C, B) +# msetsum(g, C, B)"
 apply (subgoal_tac "\<forall>x \<in> A. multiset (f(x) +# g(x)) \<and> mset_of (f(x) +# g(x)) \<subseteq> B")
 apply (erule Fin_induct)
 apply (auto simp add: munion_ac) 
--- a/src/ZF/UNITY/Mutex.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Mutex.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -36,37 +36,37 @@
 
 definition
   (** The program for process U **)
-  "U0 \<equiv> {<s,t>:state*state. t = s(u:=1, m:=#1) \<and> s`m = #0}"
+  "U0 \<equiv> {\<langle>s,t\<rangle>:state*state. t = s(u:=1, m:=#1) \<and> s`m = #0}"
 
 definition
-  "U1 \<equiv> {<s,t>:state*state. t = s(p:= s`v, m:=#2) \<and>  s`m = #1}"
+  "U1 \<equiv> {\<langle>s,t\<rangle>:state*state. t = s(p:= s`v, m:=#2) \<and>  s`m = #1}"
 
 definition
-  "U2 \<equiv> {<s,t>:state*state. t = s(m:=#3) \<and> s`p=0 \<and> s`m = #2}"
+  "U2 \<equiv> {\<langle>s,t\<rangle>:state*state. t = s(m:=#3) \<and> s`p=0 \<and> s`m = #2}"
 
 definition
-  "U3 \<equiv> {<s,t>:state*state. t=s(u:=0, m:=#4) \<and> s`m = #3}"
+  "U3 \<equiv> {\<langle>s,t\<rangle>:state*state. t=s(u:=0, m:=#4) \<and> s`m = #3}"
 
 definition
-  "U4 \<equiv> {<s,t>:state*state. t = s(p:=1, m:=#0) \<and> s`m = #4}"
+  "U4 \<equiv> {\<langle>s,t\<rangle>:state*state. t = s(p:=1, m:=#0) \<and> s`m = #4}"
 
 
    (** The program for process V **)
 
 definition
-  "V0 \<equiv> {<s,t>:state*state. t = s (v:=1, n:=#1) \<and> s`n = #0}"
+  "V0 \<equiv> {\<langle>s,t\<rangle>:state*state. t = s (v:=1, n:=#1) \<and> s`n = #0}"
 
 definition
-  "V1 \<equiv> {<s,t>:state*state. t = s(p:=not(s`u), n:=#2) \<and> s`n = #1}"
+  "V1 \<equiv> {\<langle>s,t\<rangle>:state*state. t = s(p:=not(s`u), n:=#2) \<and> s`n = #1}"
 
 definition
-  "V2 \<equiv> {<s,t>:state*state. t  = s(n:=#3) \<and> s`p=1 \<and> s`n = #2}"
+  "V2 \<equiv> {\<langle>s,t\<rangle>:state*state. t  = s(n:=#3) \<and> s`p=1 \<and> s`n = #2}"
 
 definition
-  "V3 \<equiv> {<s,t>:state*state. t = s (v:=0, n:=#4) \<and> s`n = #3}"
+  "V3 \<equiv> {\<langle>s,t\<rangle>:state*state. t = s (v:=0, n:=#4) \<and> s`n = #3}"
 
 definition
-  "V4 \<equiv> {<s,t>:state*state. t  = s (p:=0, n:=#0) \<and> s`n = #4}"
+  "V4 \<equiv> {\<langle>s,t\<rangle>:state*state. t  = s (p:=0, n:=#0) \<and> s`n = #4}"
 
 definition
   "Mutex \<equiv> mk_program({s:state. s`u=0 \<and> s`v=0 \<and> s`m = #0 \<and> s`n = #0},
--- a/src/ZF/UNITY/State.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/State.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -17,8 +17,8 @@
   type_intros  nat_subset_univ [THEN list_subset_univ, THEN subsetD]
 
 consts
-  type_of :: "i=>i"
-  default_val :: "i=>i"
+  type_of :: "i\<Rightarrow>i"
+  default_val :: "i\<Rightarrow>i"
 
 definition
   "state \<equiv> \<Prod>x \<in> var. cons(default_val(x), type_of(x))"
@@ -27,13 +27,13 @@
   "st0 \<equiv> \<lambda>x \<in> var. default_val(x)"
   
 definition
-  st_set  :: "i=>o"  where
+  st_set  :: "i\<Rightarrow>o"  where
 (* To prevent typing conditions like `A<=state' from
    being used in combination with the rules `constrains_weaken', etc. *)
   "st_set(A) \<equiv> A<=state"
 
 definition
-  st_compl :: "i=>i"  where
+  st_compl :: "i\<Rightarrow>i"  where
   "st_compl(A) \<equiv> state-A"
 
 
--- a/src/ZF/UNITY/SubstAx.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/SubstAx.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -13,11 +13,11 @@
 
 definition
   (* The definitions below are not `conventional', but yield simpler rules *)
-  Ensures :: "[i,i] => i"            (infixl \<open>Ensures\<close> 60)  where
+  Ensures :: "[i,i] \<Rightarrow> i"            (infixl \<open>Ensures\<close> 60)  where
   "A Ensures B \<equiv> {F \<in> program. F \<in> (reachable(F) \<inter> A) ensures (reachable(F) \<inter> B) }"
 
 definition
-  LeadsTo :: "[i, i] => i"            (infixl \<open>\<longmapsto>w\<close> 60)  where
+  LeadsTo :: "[i, i] \<Rightarrow> i"            (infixl \<open>\<longmapsto>w\<close> 60)  where
   "A \<longmapsto>w B \<equiv> {F \<in> program. F:(reachable(F) \<inter> A) \<longmapsto> (reachable(F) \<inter> B)}"
 
 
@@ -280,7 +280,7 @@
 
 lemma LessThan_induct: "\<lbrakk>\<forall>m \<in> nat. F:(A \<inter> f-``{m}) \<longmapsto>w ((A \<inter> f-``m) \<union> B);
       A<=f-``nat; F \<in> program\<rbrakk> \<Longrightarrow> F \<in> A \<longmapsto>w B"
-apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN LeadsTo_wf_induct])
+apply (rule_tac A1 = nat and f1 = "\<lambda>x. x" in wf_measure [THEN LeadsTo_wf_induct])
 apply (simp_all add: nat_measure_field)
 apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric])
 done
--- a/src/ZF/UNITY/UNITY.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/UNITY.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -19,7 +19,7 @@
                id(state) \<in> acts \<and> id(state) \<in> allowed}"
 
 definition
-  mk_program :: "[i,i,i]=>i"  where
+  mk_program :: "[i,i,i]\<Rightarrow>i"  where
   \<comment> \<open>The definition yields a program thanks to the coercions
        init \<inter> state, acts \<inter> Pow(state*state), etc.\<close>
   "mk_program(init, acts, allowed) \<equiv>
@@ -32,69 +32,69 @@
 
   (* Coercion from anything to program *)
 definition
-  programify :: "i=>i"  where
+  programify :: "i\<Rightarrow>i"  where
   "programify(F) \<equiv> if F \<in> program then F else SKIP"
 
 definition
-  RawInit :: "i=>i"  where
+  RawInit :: "i\<Rightarrow>i"  where
   "RawInit(F) \<equiv> fst(F)"
 
 definition
-  Init :: "i=>i"  where
+  Init :: "i\<Rightarrow>i"  where
   "Init(F) \<equiv> RawInit(programify(F))"
 
 definition
-  RawActs :: "i=>i"  where
+  RawActs :: "i\<Rightarrow>i"  where
   "RawActs(F) \<equiv> cons(id(state), fst(snd(F)))"
 
 definition
-  Acts :: "i=>i"  where
+  Acts :: "i\<Rightarrow>i"  where
   "Acts(F) \<equiv> RawActs(programify(F))"
 
 definition
-  RawAllowedActs :: "i=>i"  where
+  RawAllowedActs :: "i\<Rightarrow>i"  where
   "RawAllowedActs(F) \<equiv> cons(id(state), snd(snd(F)))"
 
 definition
-  AllowedActs :: "i=>i"  where
+  AllowedActs :: "i\<Rightarrow>i"  where
   "AllowedActs(F) \<equiv> RawAllowedActs(programify(F))"
 
 
 definition
-  Allowed :: "i =>i"  where
+  Allowed :: "i \<Rightarrow>i"  where
   "Allowed(F) \<equiv> {G \<in> program. Acts(G) \<subseteq> AllowedActs(F)}"
 
 definition
-  initially :: "i=>i"  where
+  initially :: "i\<Rightarrow>i"  where
   "initially(A) \<equiv> {F \<in> program. Init(F)\<subseteq>A}"
 
-definition "constrains" :: "[i, i] => i"  (infixl \<open>co\<close> 60)  where
+definition "constrains" :: "[i, i] \<Rightarrow> i"  (infixl \<open>co\<close> 60)  where
   "A co B \<equiv> {F \<in> program. (\<forall>act \<in> Acts(F). act``A\<subseteq>B) \<and> st_set(A)}"
     \<comment> \<open>the condition \<^term>\<open>st_set(A)\<close> makes the definition slightly
          stronger than the HOL one\<close>
 
-definition unless :: "[i, i] => i"  (infixl \<open>unless\<close> 60)  where
+definition unless :: "[i, i] \<Rightarrow> i"  (infixl \<open>unless\<close> 60)  where
   "A unless B \<equiv> (A - B) co (A \<union> B)"
 
 definition
-  stable     :: "i=>i"  where
+  stable     :: "i\<Rightarrow>i"  where
    "stable(A) \<equiv> A co A"
 
 definition
-  strongest_rhs :: "[i, i] => i"  where
+  strongest_rhs :: "[i, i] \<Rightarrow> i"  where
   "strongest_rhs(F, A) \<equiv> \<Inter>({B \<in> Pow(state). F \<in> A co B})"
 
 definition
-  invariant :: "i => i"  where
+  invariant :: "i \<Rightarrow> i"  where
   "invariant(A) \<equiv> initially(A) \<inter> stable(A)"
 
   (* meta-function composition *)
 definition
-  metacomp :: "[i=>i, i=>i] => (i=>i)" (infixl \<open>comp\<close> 65)  where
-  "f comp g \<equiv> %x. f(g(x))"
+  metacomp :: "[i\<Rightarrow>i, i\<Rightarrow>i] \<Rightarrow> (i\<Rightarrow>i)" (infixl \<open>comp\<close> 65)  where
+  "f comp g \<equiv> \<lambda>x. f(g(x))"
 
 definition
-  pg_compl :: "i=>i"  where
+  pg_compl :: "i\<Rightarrow>i"  where
   "pg_compl(X)\<equiv> program - X"
 
 text\<open>SKIP\<close>
--- a/src/ZF/UNITY/Union.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/Union.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -15,34 +15,34 @@
 
 definition
   (*FIXME: conjoin Init(F) \<inter> Init(G) \<noteq> 0 *)
-  ok :: "[i, i] => o"     (infixl \<open>ok\<close> 65)  where
+  ok :: "[i, i] \<Rightarrow> o"     (infixl \<open>ok\<close> 65)  where
     "F ok G \<equiv> Acts(F) \<subseteq> AllowedActs(G) \<and>
                Acts(G) \<subseteq> AllowedActs(F)"
 
 definition
   (*FIXME: conjoin (\<Inter>i \<in> I. Init(F(i))) \<noteq> 0 *)
-  OK  :: "[i, i=>i] => o"  where
+  OK  :: "[i, i\<Rightarrow>i] \<Rightarrow> o"  where
     "OK(I,F) \<equiv> (\<forall>i \<in> I. \<forall>j \<in> I-{i}. Acts(F(i)) \<subseteq> AllowedActs(F(j)))"
 
 definition
-  JOIN  :: "[i, i=>i] => i"  where
+  JOIN  :: "[i, i\<Rightarrow>i] \<Rightarrow> i"  where
   "JOIN(I,F) \<equiv> if I = 0 then SKIP else
                  mk_program(\<Inter>i \<in> I. Init(F(i)), \<Union>i \<in> I. Acts(F(i)),
                  \<Inter>i \<in> I. AllowedActs(F(i)))"
 
 definition
-  Join :: "[i, i] => i"  (infixl \<open>\<squnion>\<close> 65)  where
+  Join :: "[i, i] \<Rightarrow> i"  (infixl \<open>\<squnion>\<close> 65)  where
   "F \<squnion> G \<equiv> mk_program (Init(F) \<inter> Init(G), Acts(F) \<union> Acts(G),
                                 AllowedActs(F) \<inter> AllowedActs(G))"
 definition
   (*Characterizes safety properties.  Used with specifying AllowedActs*)
-  safety_prop :: "i => o"  where
+  safety_prop :: "i \<Rightarrow> o"  where
   "safety_prop(X) \<equiv> X\<subseteq>program \<and>
       SKIP \<in> X \<and> (\<forall>G \<in> program. Acts(G) \<subseteq> (\<Union>F \<in> X. Acts(F)) \<longrightarrow> G \<in> X)"
 
 syntax
-  "_JOIN1"  :: "[pttrns, i] => i"     (\<open>(3\<Squnion>_./ _)\<close> 10)
-  "_JOIN"   :: "[pttrn, i, i] => i"   (\<open>(3\<Squnion>_ \<in> _./ _)\<close> 10)
+  "_JOIN1"  :: "[pttrns, i] \<Rightarrow> i"     (\<open>(3\<Squnion>_./ _)\<close> 10)
+  "_JOIN"   :: "[pttrn, i, i] \<Rightarrow> i"   (\<open>(3\<Squnion>_ \<in> _./ _)\<close> 10)
 
 translations
   "\<Squnion>x \<in> A. B" == "CONST JOIN(A, (\<lambda>x. B))"
@@ -141,10 +141,10 @@
 
 subsection\<open>Eliminating programify form JOIN and OK expressions\<close>
 
-lemma OK_programify [iff]: "OK(I, %x. programify(F(x))) \<longleftrightarrow> OK(I, F)"
+lemma OK_programify [iff]: "OK(I, \<lambda>x. programify(F(x))) \<longleftrightarrow> OK(I, F)"
 by (simp add: OK_def)
 
-lemma JOIN_programify [iff]: "JOIN(I, %x. programify(F(x))) = JOIN(I, F)"
+lemma JOIN_programify [iff]: "JOIN(I, \<lambda>x. programify(F(x))) = JOIN(I, F)"
 by (simp add: JOIN_def)
 
 
@@ -422,7 +422,7 @@
 
 lemmas ok_sym = ok_commute [THEN iffD1]
 
-lemma ok_iff_OK: "OK({<0,F>,<1,G>,<2,H>}, snd) \<longleftrightarrow> (F ok G \<and> (F \<squnion> G) ok H)"
+lemma ok_iff_OK: "OK({\<langle>0,F\<rangle>,\<langle>1,G\<rangle>,\<langle>2,H\<rangle>}, snd) \<longleftrightarrow> (F ok G \<and> (F \<squnion> G) ok H)"
 by (simp add: ok_def Join_def OK_def Int_assoc cons_absorb
                  Int_Un_distrib2 Ball_def,  safe, force+)
 
--- a/src/ZF/UNITY/WFair.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/UNITY/WFair.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -16,26 +16,26 @@
 definition
   (* This definition specifies weak fairness.  The rest of the theory
     is generic to all forms of fairness.*)
-  transient :: "i=>i"  where
+  transient :: "i\<Rightarrow>i"  where
   "transient(A) \<equiv>{F \<in> program. (\<exists>act\<in>Acts(F). A<=domain(act) \<and>
                                act``A \<subseteq> state-A) \<and> st_set(A)}"
 
 definition
-  ensures :: "[i,i] => i"       (infixl \<open>ensures\<close> 60)  where
+  ensures :: "[i,i] \<Rightarrow> i"       (infixl \<open>ensures\<close> 60)  where
   "A ensures B \<equiv> ((A-B) co (A \<union> B)) \<inter> transient(A-B)"
 
 consts
 
   (*LEADS-TO constant for the inductive definition*)
-  leads :: "[i, i]=>i"
+  leads :: "[i, i]\<Rightarrow>i"
 
 inductive
   domains
      "leads(D, F)" \<subseteq> "Pow(D)*Pow(D)"
   intros
-    Basis:  "\<lbrakk>F \<in> A ensures B;  A \<in> Pow(D); B \<in> Pow(D)\<rbrakk> \<Longrightarrow> <A,B>:leads(D, F)"
-    Trans:  "\<lbrakk><A,B> \<in> leads(D, F); <B,C> \<in> leads(D, F)\<rbrakk> \<Longrightarrow>  <A,C>:leads(D, F)"
-    Union:   "\<lbrakk>S \<in> Pow({A \<in> S. <A, B>:leads(D, F)}); B \<in> Pow(D); S \<in> Pow(Pow(D))\<rbrakk> \<Longrightarrow>
+    Basis:  "\<lbrakk>F \<in> A ensures B;  A \<in> Pow(D); B \<in> Pow(D)\<rbrakk> \<Longrightarrow> \<langle>A,B\<rangle>:leads(D, F)"
+    Trans:  "\<lbrakk>\<langle>A,B\<rangle> \<in> leads(D, F); \<langle>B,C\<rangle> \<in> leads(D, F)\<rbrakk> \<Longrightarrow>  \<langle>A,C\<rangle>:leads(D, F)"
+    Union:   "\<lbrakk>S \<in> Pow({A \<in> S. \<langle>A, B\<rangle>:leads(D, F)}); B \<in> Pow(D); S \<in> Pow(Pow(D))\<rbrakk> \<Longrightarrow>
               <\<Union>(S),B>:leads(D, F)"
 
   monos        Pow_mono
@@ -43,12 +43,12 @@
 
 definition
   (* The Visible version of the LEADS-TO relation*)
-  leadsTo :: "[i, i] => i"       (infixl \<open>\<longmapsto>\<close> 60)  where
-  "A \<longmapsto> B \<equiv> {F \<in> program. <A,B>:leads(state, F)}"
+  leadsTo :: "[i, i] \<Rightarrow> i"       (infixl \<open>\<longmapsto>\<close> 60)  where
+  "A \<longmapsto> B \<equiv> {F \<in> program. \<langle>A,B\<rangle>:leads(state, F)}"
 
 definition
   (* wlt(F, B) is the largest set that leads to B*)
-  wlt :: "[i, i] => i"  where
+  wlt :: "[i, i] \<Rightarrow> i"  where
     "wlt(F, B) \<equiv> \<Union>({A \<in> Pow(state). F \<in> A \<longmapsto> B})"
 
 (** Ad-hoc set-theory rules **)
@@ -510,7 +510,7 @@
  apply (erule_tac I = I in leadsTo_wf_induct_aux, assumption+, best)
 done
 
-lemma nat_measure_field: "field(measure(nat, %x. x)) = nat"
+lemma nat_measure_field: "field(measure(nat, \<lambda>x. x)) = nat"
 apply (unfold field_def)
 apply (simp add: measure_def)
 apply (rule equalityI, force, clarify)
@@ -522,7 +522,7 @@
 done
 
 
-lemma Image_inverse_lessThan: "k<A \<Longrightarrow> measure(A, %x. x) -`` {k} = k"
+lemma Image_inverse_lessThan: "k<A \<Longrightarrow> measure(A, \<lambda>x. x) -`` {k} = k"
 apply (rule equalityI)
 apply (auto simp add: measure_def)
 apply (blast intro: ltD)
@@ -538,7 +538,7 @@
      F \<in> program; st_set(A); st_set(B);
      \<forall>m \<in> nat. F:(A \<inter> f-``{m}) \<longmapsto> ((A \<inter> f -`` m) \<union> B)\<rbrakk>
       \<Longrightarrow> F \<in> A \<longmapsto> B"
-apply (rule_tac A1 = nat and f1 = "%x. x" in wf_measure [THEN leadsTo_wf_induct])
+apply (rule_tac A1 = nat and f1 = "\<lambda>x. x" in wf_measure [THEN leadsTo_wf_induct])
 apply (simp_all add: nat_measure_field)
 apply (simp add: ltI Image_inverse_lessThan vimage_def [symmetric])
 done
@@ -605,7 +605,7 @@
  apply (blast intro: leadsTo_123_aux leadsTo_Un_Un leadsTo_cancel1 leadsTo_Un_duplicate, blast)
 txt\<open>Union\<close>
 apply (clarify dest!: ball_conj_distrib [THEN iffD1])
-apply (subgoal_tac "\<exists>y. y \<in> Pi (S, %A. {Ba \<in> Pow (state) . A<=Ba \<and> F \<in> Ba \<longmapsto> B \<and> F \<in> Ba - B co Ba \<union> B}) ")
+apply (subgoal_tac "\<exists>y. y \<in> Pi (S, \<lambda>A. {Ba \<in> Pow (state) . A<=Ba \<and> F \<in> Ba \<longmapsto> B \<and> F \<in> Ba - B co Ba \<union> B}) ")
 defer 1
 apply (rule AC_ball_Pi, safe)
 apply (rotate_tac 1)
--- a/src/ZF/Univ.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Univ.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -14,26 +14,26 @@
 theory Univ imports Epsilon Cardinal begin
 
 definition
-  Vfrom       :: "[i,i]=>i"  where
-    "Vfrom(A,i) \<equiv> transrec(i, %x f. A \<union> (\<Union>y\<in>x. Pow(f`y)))"
+  Vfrom       :: "[i,i]\<Rightarrow>i"  where
+    "Vfrom(A,i) \<equiv> transrec(i, \<lambda>x f. A \<union> (\<Union>y\<in>x. Pow(f`y)))"
 
 abbreviation
-  Vset :: "i=>i" where
+  Vset :: "i\<Rightarrow>i" where
   "Vset(x) \<equiv> Vfrom(0,x)"
 
 
 definition
-  Vrec        :: "[i, [i,i]=>i] =>i"  where
-    "Vrec(a,H) \<equiv> transrec(rank(a), %x g. \<lambda>z\<in>Vset(succ(x)).
+  Vrec        :: "[i, [i,i]\<Rightarrow>i] \<Rightarrow>i"  where
+    "Vrec(a,H) \<equiv> transrec(rank(a), \<lambda>x g. \<lambda>z\<in>Vset(succ(x)).
                            H(z, \<lambda>w\<in>Vset(x). g`rank(w)`w)) ` a"
 
 definition
-  Vrecursor   :: "[[i,i]=>i, i] =>i"  where
-    "Vrecursor(H,a) \<equiv> transrec(rank(a), %x g. \<lambda>z\<in>Vset(succ(x)).
+  Vrecursor   :: "[[i,i]\<Rightarrow>i, i] \<Rightarrow>i"  where
+    "Vrecursor(H,a) \<equiv> transrec(rank(a), \<lambda>x g. \<lambda>z\<in>Vset(succ(x)).
                                 H(\<lambda>w\<in>Vset(x). g`rank(w)`w, z)) ` a"
 
 definition
-  univ        :: "i=>i"  where
+  univ        :: "i\<Rightarrow>i"  where
     "univ(A) \<equiv> Vfrom(A,nat)"
 
 
@@ -124,7 +124,7 @@
 by (rule subset_mem_Vfrom, safe)
 
 lemma Pair_in_Vfrom:
-    "\<lbrakk>a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Vfrom(A,succ(succ(i)))"
+    "\<lbrakk>a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Vfrom(A,succ(succ(i)))"
 apply (unfold Pair_def)
 apply (blast intro: doubleton_in_Vfrom)
 done
@@ -219,7 +219,7 @@
 done
 
 lemma Pair_in_VLimit:
-    "\<lbrakk>a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i)\<rbrakk> \<Longrightarrow> <a,b> \<in> Vfrom(A,i)"
+    "\<lbrakk>a \<in> Vfrom(A,i);  b \<in> Vfrom(A,i);  Limit(i)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Vfrom(A,i)"
 txt\<open>Infer that a, b occur at ordinals x,xa < i.\<close>
 apply (erule Limit_VfromE, assumption)
 apply (erule Limit_VfromE, assumption)
@@ -283,12 +283,12 @@
 apply (erule Transset_Vfrom [THEN Transset_iff_Pow [THEN iffD1]])
 done
 
-lemma Transset_Pair_subset: "\<lbrakk><a,b> \<subseteq> C; Transset(C)\<rbrakk> \<Longrightarrow> a: C \<and> b: C"
+lemma Transset_Pair_subset: "\<lbrakk>\<langle>a,b\<rangle> \<subseteq> C; Transset(C)\<rbrakk> \<Longrightarrow> a: C \<and> b: C"
 by (unfold Pair_def Transset_def, blast)
 
 lemma Transset_Pair_subset_VLimit:
-     "\<lbrakk><a,b> \<subseteq> Vfrom(A,i);  Transset(A);  Limit(i)\<rbrakk>
-      \<Longrightarrow> <a,b> \<in> Vfrom(A,i)"
+     "\<lbrakk>\<langle>a,b\<rangle> \<subseteq> Vfrom(A,i);  Transset(A);  Limit(i)\<rbrakk>
+      \<Longrightarrow> \<langle>a,b\<rangle> \<in> Vfrom(A,i)"
 apply (erule Transset_Pair_subset [THEN conjE])
 apply (erule Transset_Vfrom)
 apply (blast intro: Pair_in_VLimit)
@@ -496,7 +496,7 @@
 apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
 done
 
-text\<open>This form avoids giant explosions in proofs.  NOTE USE OF \<equiv>\<close>
+text\<open>This form avoids giant explosions in proofs. NOTE the form of the premise!\<close>
 lemma def_Vrec:
     "\<lbrakk>\<And>x. h(x)\<equiv>Vrec(x,H)\<rbrakk> \<Longrightarrow>
      h(a) = H(a, \<lambda>x\<in>Vset(rank(a)). h(x))"
@@ -512,7 +512,7 @@
 apply (rule refl [THEN lam_cong, THEN subst_context], simp add: lt_def)
 done
 
-text\<open>This form avoids giant explosions in proofs.  NOTE USE OF \<equiv>\<close>
+text\<open>This form avoids giant explosions in proofs. NOTE the form of the premise!\<close>
 lemma def_Vrecursor:
      "h \<equiv> Vrecursor(H) \<Longrightarrow> h(a) = H(\<lambda>x\<in>Vset(rank(a)). h(x),  a)"
 apply simp
@@ -595,7 +595,7 @@
 done
 
 lemma Pair_in_univ:
-    "\<lbrakk>a: univ(A);  b: univ(A)\<rbrakk> \<Longrightarrow> <a,b> \<in> univ(A)"
+    "\<lbrakk>a: univ(A);  b: univ(A)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> univ(A)"
 apply (unfold univ_def)
 apply (blast intro: Pair_in_VLimit Limit_nat)
 done
@@ -619,8 +619,8 @@
 apply (rule i_subset_Vfrom)
 done
 
-text\<open>n:nat \<Longrightarrow> n:univ(A)\<close>
-lemmas nat_into_univ = nat_subset_univ [THEN subsetD]
+lemma nat_into_univ: "n \<in> nat \<Longrightarrow> n \<in> univ(A)"
+  by (rule nat_subset_univ [THEN subsetD])
 
 subsubsection\<open>Instances for 1 and 2\<close>
 
@@ -766,16 +766,16 @@
 text\<open>This weaker version says a, b exist at the same level\<close>
 lemmas Vfrom_doubleton_D = Transset_Vfrom [THEN Transset_doubleton_D]
 
-(** Using only the weaker theorem would prove <a,b> \<in> Vfrom(X,i)
+(** Using only the weaker theorem would prove \<langle>a,b\<rangle> \<in> Vfrom(X,i)
       implies a, b \<in> Vfrom(X,i), which is useless for induction.
-    Using only the stronger theorem would prove <a,b> \<in> Vfrom(X,succ(succ(i)))
+    Using only the stronger theorem would prove \<langle>a,b\<rangle> \<in> Vfrom(X,succ(succ(i)))
       implies a, b \<in> Vfrom(X,i), leaving the succ(i) case untreated.
     The combination gives a reduction by precisely one level, which is
       most convenient for proofs.
 **)
 
 lemma Pair_in_Vfrom_D:
-    "\<lbrakk><a,b> \<in> Vfrom(X,succ(i));  Transset(X)\<rbrakk>
+    "\<lbrakk>\<langle>a,b\<rangle> \<in> Vfrom(X,succ(i));  Transset(X)\<rbrakk>
      \<Longrightarrow> a \<in> Vfrom(X,i)  \<and>  b \<in> Vfrom(X,i)"
 apply (unfold Pair_def)
 apply (blast dest!: doubleton_in_Vfrom_D Vfrom_doubleton_D)
--- a/src/ZF/WF.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/WF.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -19,34 +19,34 @@
 theory WF imports Trancl begin
 
 definition
-  wf           :: "i=>o"  where
+  wf           :: "i\<Rightarrow>o"  where
     (*r is a well-founded relation*)
-    "wf(r) \<equiv> \<forall>Z. Z=0 | (\<exists>x\<in>Z. \<forall>y. <y,x>:r \<longrightarrow> \<not> y \<in> Z)"
+    "wf(r) \<equiv> \<forall>Z. Z=0 | (\<exists>x\<in>Z. \<forall>y. \<langle>y,x\<rangle>:r \<longrightarrow> \<not> y \<in> Z)"
 
 definition
-  wf_on        :: "[i,i]=>o"                      (\<open>wf[_]'(_')\<close>)  where
+  wf_on        :: "[i,i]\<Rightarrow>o"                      (\<open>wf[_]'(_')\<close>)  where
     (*r is well-founded on A*)
     "wf_on(A,r) \<equiv> wf(r \<inter> A*A)"
 
 definition
-  is_recfun    :: "[i, i, [i,i]=>i, i] =>o"  where
+  is_recfun    :: "[i, i, [i,i]\<Rightarrow>i, i] \<Rightarrow>o"  where
     "is_recfun(r,a,H,f) \<equiv> (f = (\<lambda>x\<in>r-``{a}. H(x, restrict(f, r-``{x}))))"
 
 definition
-  the_recfun   :: "[i, i, [i,i]=>i] =>i"  where
+  the_recfun   :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i"  where
     "the_recfun(r,a,H) \<equiv> (THE f. is_recfun(r,a,H,f))"
 
 definition
-  wftrec :: "[i, i, [i,i]=>i] =>i"  where
+  wftrec :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i"  where
     "wftrec(r,a,H) \<equiv> H(a, the_recfun(r,a,H))"
 
 definition
-  wfrec :: "[i, i, [i,i]=>i] =>i"  where
+  wfrec :: "[i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i"  where
     (*public version.  Does not require r to be transitive*)
-    "wfrec(r,a,H) \<equiv> wftrec(r^+, a, %x f. H(x, restrict(f,r-``{x})))"
+    "wfrec(r,a,H) \<equiv> wftrec(r^+, a, \<lambda>x f. H(x, restrict(f,r-``{x})))"
 
 definition
-  wfrec_on     :: "[i, i, i, [i,i]=>i] =>i"       (\<open>wfrec[_]'(_,_,_')\<close>)  where
+  wfrec_on     :: "[i, i, i, [i,i]\<Rightarrow>i] \<Rightarrow>i"       (\<open>wfrec[_]'(_,_,_')\<close>)  where
     "wfrec[A](r,a,H) \<equiv> wfrec(r \<inter> A*A, a, H)"
 
 
@@ -80,7 +80,7 @@
 text\<open>If every non-empty subset of \<^term>\<open>A\<close> has an \<^term>\<open>r\<close>-minimal element
    then we have \<^term>\<open>wf[A](r)\<close>.\<close>
 lemma wf_onI:
- assumes prem: "\<And>Z u. \<lbrakk>Z<=A;  u \<in> Z;  \<forall>x\<in>Z. \<exists>y\<in>Z. <y,x>:r\<rbrakk> \<Longrightarrow> False"
+ assumes prem: "\<And>Z u. \<lbrakk>Z<=A;  u \<in> Z;  \<forall>x\<in>Z. \<exists>y\<in>Z. \<langle>y,x\<rangle>:r\<rbrakk> \<Longrightarrow> False"
  shows         "wf[A](r)"
 apply (unfold wf_on_def wf_def)
 apply (rule equals0I [THEN disjCI, THEN allI])
@@ -89,9 +89,9 @@
 
 text\<open>If \<^term>\<open>r\<close> allows well-founded induction over \<^term>\<open>A\<close>
    then we have \<^term>\<open>wf[A](r)\<close>.   Premise is equivalent to
-  \<^prop>\<open>\<And>B. \<forall>x\<in>A. (\<forall>y. <y,x>: r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B \<Longrightarrow> A<=B\<close>\<close>
+  \<^prop>\<open>\<And>B. \<forall>x\<in>A. (\<forall>y. \<langle>y,x\<rangle>: r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B \<Longrightarrow> A<=B\<close>\<close>
 lemma wf_onI2:
- assumes prem: "\<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B;   y \<in> A\<rbrakk>
+ assumes prem: "\<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. \<langle>y,x\<rangle>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B;   y \<in> A\<rbrakk>
                        \<Longrightarrow> y \<in> B"
  shows         "wf[A](r)"
 apply (rule wf_onI)
@@ -107,7 +107,7 @@
   \<^term>\<open>P(z)\<close> does not hold...\<close>
 lemma wf_induct_raw:
     "\<lbrakk>wf(r);
-        \<And>x.\<lbrakk>\<forall>y. <y,x>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk>
+        \<And>x.\<lbrakk>\<forall>y. \<langle>y,x\<rangle>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk>
      \<Longrightarrow> P(a)"
 apply (unfold wf_def)
 apply (erule_tac x = "{z \<in> domain(r). \<not> P(z)}" in allE)
@@ -119,7 +119,7 @@
 text\<open>The form of this rule is designed to match \<open>wfI\<close>\<close>
 lemma wf_induct2:
     "\<lbrakk>wf(r);  a \<in> A;  field(r)<=A;
-        \<And>x.\<lbrakk>x \<in> A;  \<forall>y. <y,x>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk>
+        \<And>x.\<lbrakk>x \<in> A;  \<forall>y. \<langle>y,x\<rangle>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)\<rbrakk>
      \<Longrightarrow>  P(a)"
 apply (erule_tac P="a \<in> A" in rev_mp)
 apply (erule_tac a=a in wf_induct, blast)
@@ -130,7 +130,7 @@
 
 lemma wf_on_induct_raw [consumes 2, induct set: wf_on]:
     "\<lbrakk>wf[A](r);  a \<in> A;
-        \<And>x.\<lbrakk>x \<in> A;  \<forall>y\<in>A. <y,x>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)
+        \<And>x.\<lbrakk>x \<in> A;  \<forall>y\<in>A. \<langle>y,x\<rangle>: r \<longrightarrow> P(y)\<rbrakk> \<Longrightarrow> P(x)
 \<rbrakk>  \<Longrightarrow>  P(a)"
 apply (unfold wf_on_def)
 apply (erule wf_induct2, assumption)
@@ -145,7 +145,7 @@
    then we have \<^term>\<open>wf(r)\<close>.\<close>
 lemma wfI:
     "\<lbrakk>field(r)<=A;
-        \<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. <y,x>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B;  y \<in> A\<rbrakk>
+        \<And>y B. \<lbrakk>\<forall>x\<in>A. (\<forall>y\<in>A. \<langle>y,x\<rangle>:r \<longrightarrow> y \<in> B) \<longrightarrow> x \<in> B;  y \<in> A\<rbrakk>
                \<Longrightarrow> y \<in> B\<rbrakk>
      \<Longrightarrow>  wf(r)"
 apply (rule wf_on_subset_A [THEN wf_on_field_imp_wf])
@@ -157,35 +157,35 @@
 
 subsection\<open>Basic Properties of Well-Founded Relations\<close>
 
-lemma wf_not_refl: "wf(r) \<Longrightarrow> <a,a> \<notin> r"
+lemma wf_not_refl: "wf(r) \<Longrightarrow> \<langle>a,a\<rangle> \<notin> r"
 by (erule_tac a=a in wf_induct, blast)
 
-lemma wf_not_sym [rule_format]: "wf(r) \<Longrightarrow> \<forall>x. <a,x>:r \<longrightarrow> <x,a> \<notin> r"
+lemma wf_not_sym [rule_format]: "wf(r) \<Longrightarrow> \<forall>x. \<langle>a,x\<rangle>:r \<longrightarrow> \<langle>x,a\<rangle> \<notin> r"
 by (erule_tac a=a in wf_induct, blast)
 
-(* @{term"\<lbrakk>wf(r);  <a,x> \<in> r;  \<not>P \<Longrightarrow> <x,a> \<in> r\<rbrakk> \<Longrightarrow> P"} *)
+(* @{term"\<lbrakk>wf(r);  \<langle>a,x\<rangle> \<in> r;  \<not>P \<Longrightarrow> \<langle>x,a\<rangle> \<in> r\<rbrakk> \<Longrightarrow> P"} *)
 lemmas wf_asym = wf_not_sym [THEN swap]
 
-lemma wf_on_not_refl: "\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow> <a,a> \<notin> r"
+lemma wf_on_not_refl: "\<lbrakk>wf[A](r); a \<in> A\<rbrakk> \<Longrightarrow> \<langle>a,a\<rangle> \<notin> r"
 by (erule_tac a=a in wf_on_induct, assumption, blast)
 
 lemma wf_on_not_sym:
-     "\<lbrakk>wf[A](r);  a \<in> A\<rbrakk> \<Longrightarrow> (\<And>b. b\<in>A \<Longrightarrow> <a,b>:r \<Longrightarrow> <b,a>\<notin>r)"
+     "\<lbrakk>wf[A](r);  a \<in> A\<rbrakk> \<Longrightarrow> (\<And>b. b\<in>A \<Longrightarrow> \<langle>a,b\<rangle>:r \<Longrightarrow> \<langle>b,a\<rangle>\<notin>r)"
 apply (atomize (full), intro impI)
 apply (erule_tac a=a in wf_on_induct, assumption, blast)
 done
 
 lemma wf_on_asym:
-     "\<lbrakk>wf[A](r);  \<not>Z \<Longrightarrow> <a,b> \<in> r;
-         <b,a> \<notin> r \<Longrightarrow> Z; \<not>Z \<Longrightarrow> a \<in> A; \<not>Z \<Longrightarrow> b \<in> A\<rbrakk> \<Longrightarrow> Z"
+     "\<lbrakk>wf[A](r);  \<not>Z \<Longrightarrow> \<langle>a,b\<rangle> \<in> r;
+         \<langle>b,a\<rangle> \<notin> r \<Longrightarrow> Z; \<not>Z \<Longrightarrow> a \<in> A; \<not>Z \<Longrightarrow> b \<in> A\<rbrakk> \<Longrightarrow> Z"
 by (blast dest: wf_on_not_sym)
 
 
 (*Needed to prove well_ordI.  Could also reason that wf[A](r) means
   wf(r \<inter> A*A);  thus wf( (r \<inter> A*A)^+ ) and use wf_not_refl *)
 lemma wf_on_chain3:
-     "\<lbrakk>wf[A](r); <a,b>:r; <b,c>:r; <c,a>:r; a \<in> A; b \<in> A; c \<in> A\<rbrakk> \<Longrightarrow> P"
-apply (subgoal_tac "\<forall>y\<in>A. \<forall>z\<in>A. <a,y>:r \<longrightarrow> <y,z>:r \<longrightarrow> <z,a>:r \<longrightarrow> P",
+     "\<lbrakk>wf[A](r); \<langle>a,b\<rangle>:r; \<langle>b,c\<rangle>:r; \<langle>c,a\<rangle>:r; a \<in> A; b \<in> A; c \<in> A\<rbrakk> \<Longrightarrow> P"
+apply (subgoal_tac "\<forall>y\<in>A. \<forall>z\<in>A. \<langle>a,y\<rangle>:r \<longrightarrow> \<langle>y,z\<rangle>:r \<longrightarrow> \<langle>z,a\<rangle>:r \<longrightarrow> P",
        blast)
 apply (erule_tac a=a in wf_on_induct, assumption, blast)
 done
@@ -228,16 +228,16 @@
 lemmas is_recfun_imp_function = is_recfun_type [THEN fun_is_function]
 
 lemma apply_recfun:
-    "\<lbrakk>is_recfun(r,a,H,f); <x,a>:r\<rbrakk> \<Longrightarrow> f`x = H(x, restrict(f,r-``{x}))"
+    "\<lbrakk>is_recfun(r,a,H,f); \<langle>x,a\<rangle>:r\<rbrakk> \<Longrightarrow> f`x = H(x, restrict(f,r-``{x}))"
 apply (unfold is_recfun_def)
   txt\<open>replace f only on the left-hand side\<close>
-apply (erule_tac P = "%x. t(x) = u" for t u in ssubst)
+apply (erule_tac P = "\<lambda>x. t(x) = u" for t u in ssubst)
 apply (simp add: underI)
 done
 
 lemma is_recfun_equal [rule_format]:
      "\<lbrakk>wf(r);  trans(r);  is_recfun(r,a,H,f);  is_recfun(r,b,H,g)\<rbrakk>
-      \<Longrightarrow> <x,a>:r \<longrightarrow> <x,b>:r \<longrightarrow> f`x=g`x"
+      \<Longrightarrow> \<langle>x,a\<rangle>:r \<longrightarrow> \<langle>x,b\<rangle>:r \<longrightarrow> f`x=g`x"
 apply (frule_tac f = f in is_recfun_type)
 apply (frule_tac f = g in is_recfun_type)
 apply (simp add: is_recfun_def)
@@ -245,8 +245,8 @@
 apply (intro impI)
 apply (elim ssubst)
 apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def)
-apply (rule_tac t = "%z. H (x, z)" for x in subst_context)
-apply (subgoal_tac "\<forall>y\<in>r-``{x}. \<forall>z. <y,z>:f \<longleftrightarrow> <y,z>:g")
+apply (rule_tac t = "\<lambda>z. H (x, z)" for x in subst_context)
+apply (subgoal_tac "\<forall>y\<in>r-``{x}. \<forall>z. \<langle>y,z\<rangle>:f \<longleftrightarrow> \<langle>y,z\<rangle>:g")
  apply (blast dest: transD)
 apply (simp add: apply_iff)
 apply (blast dest: transD intro: sym)
@@ -254,7 +254,7 @@
 
 lemma is_recfun_cut:
      "\<lbrakk>wf(r);  trans(r);
-         is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  <b,a>:r\<rbrakk>
+         is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  \<langle>b,a\<rangle>:r\<rbrakk>
       \<Longrightarrow> restrict(f, r-``{b}) = g"
 apply (frule_tac f = f in is_recfun_type)
 apply (rule fun_extension)
@@ -292,7 +292,7 @@
 apply (rule lam_cong [OF refl])
 apply (drule underD)
 apply (fold is_recfun_def)
-apply (rule_tac t = "%z. H(x, z)" for x in subst_context)
+apply (rule_tac t = "\<lambda>z. H(x, z)" for x in subst_context)
 apply (rule fun_extension)
   apply (blast intro: is_recfun_type)
  apply (rule lam_type [THEN restrict_type2])
@@ -300,7 +300,7 @@
  apply (blast dest: transD)
 apply atomize
 apply (frule spec [THEN mp], assumption)
-apply (subgoal_tac "<xa,a1> \<in> r")
+apply (subgoal_tac "\<langle>xa,a1\<rangle> \<in> r")
  apply (drule_tac x1 = xa in spec [THEN mp], assumption)
 apply (simp add: vimage_singleton_iff
                  apply_recfun is_recfun_cut)
@@ -311,7 +311,7 @@
 subsection\<open>Unfolding \<^term>\<open>wftrec(r,a,H)\<close>\<close>
 
 lemma the_recfun_cut:
-     "\<lbrakk>wf(r);  trans(r);  <b,a>:r\<rbrakk>
+     "\<lbrakk>wf(r);  trans(r);  \<langle>b,a\<rangle>:r\<rbrakk>
       \<Longrightarrow> restrict(the_recfun(r,a,H), r-``{b}) = the_recfun(r,b,H)"
 by (blast intro: is_recfun_cut unfold_the_recfun)
 
@@ -366,7 +366,7 @@
 
 text\<open>Minimal-element characterization of well-foundedness\<close>
 lemma wf_eq_minimal:
-     "wf(r) \<longleftrightarrow> (\<forall>Q x. x \<in> Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. <y,z>:r \<longrightarrow> y\<notin>Q))"
+     "wf(r) \<longleftrightarrow> (\<forall>Q x. x \<in> Q \<longrightarrow> (\<exists>z\<in>Q. \<forall>y. \<langle>y,z\<rangle>:r \<longrightarrow> y\<notin>Q))"
 by (unfold wf_def, blast)
 
 end
--- a/src/ZF/ZF.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ZF.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -7,14 +7,14 @@
 
 subsection\<open>Iteration of the function \<^term>\<open>F\<close>\<close>
 
-consts  iterates :: "[i=>i,i,i] => i"   (\<open>(_^_ '(_'))\<close> [60,1000,1000] 60)
+consts  iterates :: "[i\<Rightarrow>i,i,i] \<Rightarrow> i"   (\<open>(_^_ '(_'))\<close> [60,1000,1000] 60)
 
 primrec
     "F^0 (x) = x"
     "F^(succ(n)) (x) = F(F^n (x))"
 
 definition
-  iterates_omega :: "[i=>i,i] => i" (\<open>(_^\<omega> '(_'))\<close> [60,1000] 60) where
+  iterates_omega :: "[i\<Rightarrow>i,i] \<Rightarrow> i" (\<open>(_^\<omega> '(_'))\<close> [60,1000] 60) where
     "F^\<omega> (x) \<equiv> \<Union>n\<in>nat. F^n (x)"
 
 lemma iterates_triv:
@@ -45,7 +45,7 @@
     three cases of ordinals\<close>
 
 definition
-  transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where
+  transrec3 :: "[i, i, [i,i]\<Rightarrow>i, [i,i]\<Rightarrow>i] \<Rightarrow>i" where
     "transrec3(k, a, b, c) \<equiv>
        transrec(k, \<lambda>x r.
          if x=0 then a
--- a/src/ZF/ZF_Base.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ZF_Base.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -49,10 +49,10 @@
    The resulting set (for functional P) is the same as with
    PrimReplace, but the rules are simpler. *)
 definition Replace :: "[i, [i, i] \<Rightarrow> o] \<Rightarrow> i"
-  where "Replace(A,P) \<equiv> PrimReplace(A, %x y. (\<exists>!z. P(x,z)) \<and> P(x,y))"
+  where "Replace(A,P) \<equiv> PrimReplace(A, \<lambda>x y. (\<exists>!z. P(x,z)) \<and> P(x,y))"
 
 syntax
-  "_Replace"  :: "[pttrn, pttrn, i, o] => i"  (\<open>(1{_ ./ _ \<in> _, _})\<close>)
+  "_Replace"  :: "[pttrn, pttrn, i, o] \<Rightarrow> i"  (\<open>(1{_ ./ _ \<in> _, _})\<close>)
 translations
   "{y. x\<in>A, Q}" \<rightleftharpoons> "CONST Replace(A, \<lambda>x y. Q)"
 
@@ -63,7 +63,7 @@
   where "RepFun(A,f) \<equiv> {y . x\<in>A, y=f(x)}"
 
 syntax
-  "_RepFun" :: "[i, pttrn, i] => i"  (\<open>(1{_ ./ _ \<in> _})\<close> [51,0,51])
+  "_RepFun" :: "[i, pttrn, i] \<Rightarrow> i"  (\<open>(1{_ ./ _ \<in> _})\<close> [51,0,51])
 translations
   "{b. x\<in>A}" \<rightleftharpoons> "CONST RepFun(A, \<lambda>x. b)"
 
@@ -81,12 +81,12 @@
 
 subsection \<open>General union and intersection\<close>
 
-definition Inter :: "i => i"  (\<open>\<Inter>_\<close> [90] 90)
+definition Inter :: "i \<Rightarrow> i"  (\<open>\<Inter>_\<close> [90] 90)
   where "\<Inter>(A) \<equiv> { x\<in>\<Union>(A) . \<forall>y\<in>A. x\<in>y}"
 
 syntax
-  "_UNION" :: "[pttrn, i, i] => i"  (\<open>(3\<Union>_\<in>_./ _)\<close> 10)
-  "_INTER" :: "[pttrn, i, i] => i"  (\<open>(3\<Inter>_\<in>_./ _)\<close> 10)
+  "_UNION" :: "[pttrn, i, i] \<Rightarrow> i"  (\<open>(3\<Union>_\<in>_./ _)\<close> 10)
+  "_INTER" :: "[pttrn, i, i] \<Rightarrow> i"  (\<open>(3\<Inter>_\<in>_./ _)\<close> 10)
 translations
   "\<Union>x\<in>A. B" == "CONST Union({B. x\<in>A})"
   "\<Inter>x\<in>A. B" == "CONST Inter({B. x\<in>A})"
@@ -97,7 +97,7 @@
 (*Unordered pairs (Upair) express binary union/intersection and cons;
   set enumerations translate as {a,...,z} = cons(a,...,cons(z,0)...)*)
 
-definition Upair :: "[i, i] => i"
+definition Upair :: "[i, i] \<Rightarrow> i"
   where "Upair(a,b) \<equiv> {y. x\<in>Pow(Pow(0)), (x=0 \<and> y=a) | (x=Pow(0) \<and> y=b)}"
 
 definition Subset :: "[i, i] \<Rightarrow> o"  (infixl \<open>\<subseteq>\<close> 50)  \<comment> \<open>subset relation\<close>
@@ -112,10 +112,10 @@
 definition Int :: "[i, i] \<Rightarrow> i"  (infixl \<open>\<inter>\<close> 70)  \<comment> \<open>binary intersection\<close>
   where "A \<inter> B \<equiv> \<Inter>(Upair(A,B))"
 
-definition cons :: "[i, i] => i"
+definition cons :: "[i, i] \<Rightarrow> i"
   where "cons(a,A) \<equiv> Upair(a,a) \<union> A"
 
-definition succ :: "i => i"
+definition succ :: "i \<Rightarrow> i"
   where "succ(i) \<equiv> cons(i, i)"
 
 nonterminal "is"
@@ -160,14 +160,14 @@
   where if_def: "if P then a else b \<equiv> THE z. P \<and> z=a | \<not>P \<and> z=b"
 
 abbreviation (input)
-  old_if :: "[o, i, i] => i"  (\<open>if '(_,_,_')\<close>)
+  old_if :: "[o, i, i] \<Rightarrow> i"  (\<open>if '(_,_,_')\<close>)
   where "if(P,a,b) \<equiv> If(P,a,b)"
 
 
 subsection \<open>Ordered Pairing\<close>
 
 (* this "symmetric" definition works better than {{a}, {a,b}} *)
-definition Pair :: "[i, i] => i"
+definition Pair :: "[i, i] \<Rightarrow> i"
   where "Pair(a,b) \<equiv> {{a,a}, {a,b}}"
 
 definition fst :: "i \<Rightarrow> i"
@@ -182,10 +182,10 @@
 (* Patterns -- extends pre-defined type "pttrn" used in abstractions *)
 nonterminal patterns
 syntax
-  "_pattern"  :: "patterns => pttrn"         (\<open>\<langle>_\<rangle>\<close>)
-  ""          :: "pttrn => patterns"         (\<open>_\<close>)
-  "_patterns" :: "[pttrn, patterns] => patterns"  (\<open>_,/_\<close>)
-  "_Tuple"    :: "[i, is] => i"              (\<open>\<langle>(_,/ _)\<rangle>\<close>)
+  "_pattern"  :: "patterns \<Rightarrow> pttrn"         (\<open>\<langle>_\<rangle>\<close>)
+  ""          :: "pttrn \<Rightarrow> patterns"         (\<open>_\<close>)
+  "_patterns" :: "[pttrn, patterns] \<Rightarrow> patterns"  (\<open>_,/_\<close>)
+  "_Tuple"    :: "[i, is] \<Rightarrow> i"              (\<open>\<langle>(_,/ _)\<rangle>\<close>)
 translations
   "\<langle>x, y, z\<rangle>"   == "\<langle>x, \<langle>y, z\<rangle>\<rangle>"
   "\<langle>x, y\<rangle>"      == "CONST Pair(x, y)"
@@ -195,7 +195,7 @@
 definition Sigma :: "[i, i \<Rightarrow> i] \<Rightarrow> i"
   where "Sigma(A,B) \<equiv> \<Union>x\<in>A. \<Union>y\<in>B(x). {\<langle>x,y\<rangle>}"
 
-abbreviation cart_prod :: "[i, i] => i"  (infixr \<open>\<times>\<close> 80)  \<comment> \<open>Cartesian product\<close>
+abbreviation cart_prod :: "[i, i] \<Rightarrow> i"  (infixr \<open>\<times>\<close> 80)  \<comment> \<open>Cartesian product\<close>
   where "A \<times> B \<equiv> Sigma(A, \<lambda>_. B)"
 
 
@@ -249,9 +249,9 @@
 (* binder syntax *)
 
 syntax
-  "_PROD"     :: "[pttrn, i, i] => i"        (\<open>(3\<Prod>_\<in>_./ _)\<close> 10)
-  "_SUM"      :: "[pttrn, i, i] => i"        (\<open>(3\<Sum>_\<in>_./ _)\<close> 10)
-  "_lam"      :: "[pttrn, i, i] => i"        (\<open>(3\<lambda>_\<in>_./ _)\<close> 10)
+  "_PROD"     :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3\<Prod>_\<in>_./ _)\<close> 10)
+  "_SUM"      :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3\<Sum>_\<in>_./ _)\<close> 10)
+  "_lam"      :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3\<lambda>_\<in>_./ _)\<close> 10)
 translations
   "\<Prod>x\<in>A. B"   == "CONST Pi(A, \<lambda>x. B)"
   "\<Sum>x\<in>A. B"   == "CONST Sigma(A, \<lambda>x. B)"
@@ -270,18 +270,18 @@
   not_mem         (infixl \<open>\<not>:\<close> 50)
 
 syntax (ASCII)
-  "_Ball"     :: "[pttrn, i, o] => o"        (\<open>(3ALL _:_./ _)\<close> 10)
-  "_Bex"      :: "[pttrn, i, o] => o"        (\<open>(3EX _:_./ _)\<close> 10)
-  "_Collect"  :: "[pttrn, i, o] => i"        (\<open>(1{_: _ ./ _})\<close>)
-  "_Replace"  :: "[pttrn, pttrn, i, o] => i" (\<open>(1{_ ./ _: _, _})\<close>)
-  "_RepFun"   :: "[i, pttrn, i] => i"        (\<open>(1{_ ./ _: _})\<close> [51,0,51])
-  "_UNION"    :: "[pttrn, i, i] => i"        (\<open>(3UN _:_./ _)\<close> 10)
-  "_INTER"    :: "[pttrn, i, i] => i"        (\<open>(3INT _:_./ _)\<close> 10)
-  "_PROD"     :: "[pttrn, i, i] => i"        (\<open>(3PROD _:_./ _)\<close> 10)
-  "_SUM"      :: "[pttrn, i, i] => i"        (\<open>(3SUM _:_./ _)\<close> 10)
-  "_lam"      :: "[pttrn, i, i] => i"        (\<open>(3lam _:_./ _)\<close> 10)
-  "_Tuple"    :: "[i, is] => i"              (\<open><(_,/ _)>\<close>)
-  "_pattern"  :: "patterns => pttrn"         (\<open><_>\<close>)
+  "_Ball"     :: "[pttrn, i, o] \<Rightarrow> o"        (\<open>(3ALL _:_./ _)\<close> 10)
+  "_Bex"      :: "[pttrn, i, o] \<Rightarrow> o"        (\<open>(3EX _:_./ _)\<close> 10)
+  "_Collect"  :: "[pttrn, i, o] \<Rightarrow> i"        (\<open>(1{_: _ ./ _})\<close>)
+  "_Replace"  :: "[pttrn, pttrn, i, o] \<Rightarrow> i" (\<open>(1{_ ./ _: _, _})\<close>)
+  "_RepFun"   :: "[i, pttrn, i] \<Rightarrow> i"        (\<open>(1{_ ./ _: _})\<close> [51,0,51])
+  "_UNION"    :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3UN _:_./ _)\<close> 10)
+  "_INTER"    :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3INT _:_./ _)\<close> 10)
+  "_PROD"     :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3PROD _:_./ _)\<close> 10)
+  "_SUM"      :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3SUM _:_./ _)\<close> 10)
+  "_lam"      :: "[pttrn, i, i] \<Rightarrow> i"        (\<open>(3lam _:_./ _)\<close> 10)
+  "_Tuple"    :: "[i, is] \<Rightarrow> i"              (\<open><(_,/ _)>\<close>)
+  "_pattern"  :: "patterns \<Rightarrow> pttrn"         (\<open><_>\<close>)
 
 
 subsection \<open>Substitution\<close>
@@ -510,7 +510,7 @@
 
 lemma Collect_cong [cong]:
     "\<lbrakk>A=B;  \<And>x. x\<in>B \<Longrightarrow> P(x) <-> Q(x)\<rbrakk>
-     \<Longrightarrow> Collect(A, %x. P(x)) = Collect(B, %x. Q(x))"
+     \<Longrightarrow> Collect(A, \<lambda>x. P(x)) = Collect(B, \<lambda>x. Q(x))"
 by (simp add: Collect_def)
 
 
--- a/src/ZF/Zorn.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/Zorn.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -11,23 +11,23 @@
 Proofs of Some Classical Theorems of Set Theory,'' by Abrial and Laffitte.\<close>
 
 definition
-  Subset_rel :: "i=>i"  where
-   "Subset_rel(A) \<equiv> {z \<in> A*A . \<exists>x y. z=<x,y> \<and> x<=y \<and> x\<noteq>y}"
+  Subset_rel :: "i\<Rightarrow>i"  where
+   "Subset_rel(A) \<equiv> {z \<in> A*A . \<exists>x y. z=\<langle>x,y\<rangle> \<and> x<=y \<and> x\<noteq>y}"
 
 definition
-  chain      :: "i=>i"  where
+  chain      :: "i\<Rightarrow>i"  where
    "chain(A)      \<equiv> {F \<in> Pow(A). \<forall>X\<in>F. \<forall>Y\<in>F. X<=Y | Y<=X}"
 
 definition
-  super      :: "[i,i]=>i"  where
+  super      :: "[i,i]\<Rightarrow>i"  where
    "super(A,c)    \<equiv> {d \<in> chain(A). c<=d \<and> c\<noteq>d}"
 
 definition
-  maxchain   :: "i=>i"  where
+  maxchain   :: "i\<Rightarrow>i"  where
    "maxchain(A)   \<equiv> {c \<in> chain(A). super(A,c)=0}"
 
 definition
-  increasing :: "i=>i"  where
+  increasing :: "i\<Rightarrow>i"  where
     "increasing(A) \<equiv> {f \<in> Pow(A)->Pow(A). \<forall>x. x<=A \<longrightarrow> x<=f`x}"
 
 
@@ -42,7 +42,7 @@
     the induction rule lets us assume that condition!  Many inductive proofs
     are therefore unconditional.\<close>
 consts
-  "TFin" :: "[i,i]=>i"
+  "TFin" :: "[i,i]\<Rightarrow>i"
 
 inductive
   domains       "TFin(S,next)" \<subseteq> "Pow(S)"
@@ -396,7 +396,7 @@
          \<forall>X \<in> Pow(S). next`X = if(X=S, S, cons(ch`(S-X), X))\<rbrakk>
       \<Longrightarrow> (\<lambda> x \<in> S. \<Union>({y \<in> TFin(S,next). x \<notin> y}))
                \<in> inj(S, TFin(S,next) - {S})"
-apply (rule_tac d = "%y. ch` (S-y) " in lam_injective)
+apply (rule_tac d = "\<lambda>y. ch` (S-y) " in lam_injective)
 apply (rule DiffI)
 apply (rule Collect_subset [THEN TFin_UnionI])
 apply (blast intro!: Collect_subset [THEN TFin_UnionI] elim: equalityE)
@@ -435,8 +435,8 @@
 text \<open>Reimported from HOL by Clemens Ballarin.\<close>
 
 
-definition Chain :: "i => i" where
-  "Chain(r) = {A \<in> Pow(field(r)). \<forall>a\<in>A. \<forall>b\<in>A. <a, b> \<in> r | <b, a> \<in> r}"
+definition Chain :: "i \<Rightarrow> i" where
+  "Chain(r) = {A \<in> Pow(field(r)). \<forall>a\<in>A. \<forall>b\<in>A. \<langle>a, b\<rangle> \<in> r | \<langle>b, a\<rangle> \<in> r}"
 
 lemma mono_Chain:
   "r \<subseteq> s \<Longrightarrow> Chain(r) \<subseteq> Chain(s)"
@@ -445,8 +445,8 @@
 
 theorem Zorn_po:
   assumes po: "Partial_order(r)"
-    and u: "\<forall>C\<in>Chain(r). \<exists>u\<in>field(r). \<forall>a\<in>C. <a, u> \<in> r"
-  shows "\<exists>m\<in>field(r). \<forall>a\<in>field(r). <m, a> \<in> r \<longrightarrow> a = m"
+    and u: "\<forall>C\<in>Chain(r). \<exists>u\<in>field(r). \<forall>a\<in>C. \<langle>a, u\<rangle> \<in> r"
+  shows "\<exists>m\<in>field(r). \<forall>a\<in>field(r). \<langle>m, a\<rangle> \<in> r \<longrightarrow> a = m"
 proof -
   have "Preorder(r)" using po by (simp add: partial_order_on_def)
   \<comment> \<open>Mirror r in the set of subsets below (wrt r) elements of A (?).\<close>
@@ -477,11 +477,11 @@
       fix a b
       assume "a \<in> field(r)" "r -`` {a} \<in> C" "b \<in> field(r)" "r -`` {b} \<in> C"
       hence "r -`` {a} \<subseteq> r -`` {b} | r -`` {b} \<subseteq> r -`` {a}" using 2 by auto
-      then show "<a, b> \<in> r | <b, a> \<in> r"
+      then show "\<langle>a, b\<rangle> \<in> r | \<langle>b, a\<rangle> \<in> r"
         using \<open>Preorder(r)\<close> \<open>a \<in> field(r)\<close> \<open>b \<in> field(r)\<close>
         by (simp add: subset_vimage1_vimage1_iff)
     qed
-    then obtain u where uA: "u \<in> field(r)" "\<forall>a\<in>?A. <a, u> \<in> r"
+    then obtain u where uA: "u \<in> field(r)" "\<forall>a\<in>?A. \<langle>a, u\<rangle> \<in> r"
       using u
       apply auto
       apply (drule bspec) apply assumption
@@ -498,7 +498,7 @@
         apply (erule lamE)
         apply simp
         done
-      then show "<a, u> \<in> r" using uA aB \<open>Preorder(r)\<close>
+      then show "\<langle>a, u\<rangle> \<in> r" using uA aB \<open>Preorder(r)\<close>
         by (auto simp: preorder_on_def refl_def) (blast dest: trans_onD)+
     qed
     then show "\<exists>U\<in>field(r). \<forall>A\<in>C. A \<subseteq> r -`` {U}"
@@ -508,7 +508,7 @@
   obtain m B where "m \<in> field(r)" "B = r -`` {m}"
     "\<forall>x\<in>field(r). B \<subseteq> r -`` {x} \<longrightarrow> B = r -`` {x}"
     by (auto elim!: lamE simp: ball_image_simp)
-  then have "\<forall>a\<in>field(r). <m, a> \<in> r \<longrightarrow> a = m"
+  then have "\<forall>a\<in>field(r). \<langle>m, a\<rangle> \<in> r \<longrightarrow> a = m"
     using po \<open>Preorder(r)\<close> \<open>m \<in> field(r)\<close>
     by (auto simp: subset_vimage1_vimage1_iff Partial_order_eq_vimage1_vimage1_iff)
   then show ?thesis using \<open>m \<in> field(r)\<close> by blast
--- a/src/ZF/equalities.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/equalities.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -36,18 +36,18 @@
 
 subsection\<open>Converse of a Relation\<close>
 
-lemma converse_iff [simp]: "<a,b>\<in> converse(r) \<longleftrightarrow> <b,a>\<in>r"
+lemma converse_iff [simp]: "\<langle>a,b\<rangle>\<in> converse(r) \<longleftrightarrow> \<langle>b,a\<rangle>\<in>r"
 by (unfold converse_def, blast)
 
-lemma converseI [intro!]: "<a,b>\<in>r \<Longrightarrow> <b,a>\<in>converse(r)"
+lemma converseI [intro!]: "\<langle>a,b\<rangle>\<in>r \<Longrightarrow> \<langle>b,a\<rangle>\<in>converse(r)"
 by (unfold converse_def, blast)
 
-lemma converseD: "<a,b> \<in> converse(r) \<Longrightarrow> <b,a> \<in> r"
+lemma converseD: "\<langle>a,b\<rangle> \<in> converse(r) \<Longrightarrow> \<langle>b,a\<rangle> \<in> r"
 by (unfold converse_def, blast)
 
 lemma converseE [elim!]:
     "\<lbrakk>yx \<in> converse(r);
-        \<And>x y. \<lbrakk>yx=<y,x>;  <x,y>\<in>r\<rbrakk> \<Longrightarrow> P\<rbrakk>
+        \<And>x y. \<lbrakk>yx=\<langle>y,x\<rangle>;  \<langle>x,y\<rangle>\<in>r\<rbrakk> \<Longrightarrow> P\<rbrakk>
      \<Longrightarrow> P"
 by (unfold converse_def, blast)
 
@@ -482,7 +482,7 @@
 lemma INT_Union_eq:
      "0 \<notin> A \<Longrightarrow> (\<Inter>x\<in> \<Union>(A). B(x)) = (\<Inter>y\<in>A. \<Inter>x\<in>y. B(x))"
 apply (subgoal_tac "\<forall>x\<in>A. x\<noteq>0")
- prefer 2 apply blast
+   prefer 2 apply blast
 apply (force simp add: Inter_def ball_conj_distrib)
 done
 
@@ -580,14 +580,14 @@
 
 (** Domain **)
 
-lemma domain_iff: "a: domain(r) \<longleftrightarrow> (\<exists>y. <a,y>\<in> r)"
+lemma domain_iff: "a: domain(r) \<longleftrightarrow> (\<exists>y. \<langle>a,y\<rangle>\<in> r)"
 by (unfold domain_def, blast)
 
-lemma domainI [intro]: "<a,b>\<in> r \<Longrightarrow> a: domain(r)"
+lemma domainI [intro]: "\<langle>a,b\<rangle>\<in> r \<Longrightarrow> a: domain(r)"
 by (unfold domain_def, blast)
 
 lemma domainE [elim!]:
-    "\<lbrakk>a \<in> domain(r);  \<And>y. <a,y>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+    "\<lbrakk>a \<in> domain(r);  \<And>y. \<langle>a,y\<rangle>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (unfold domain_def, blast)
 
 lemma domain_subset: "domain(Sigma(A,B)) \<subseteq> A"
@@ -599,7 +599,7 @@
 lemma domain_0 [simp]: "domain(0) = 0"
 by blast
 
-lemma domain_cons [simp]: "domain(cons(<a,b>,r)) = cons(a, domain(r))"
+lemma domain_cons [simp]: "domain(cons(\<langle>a,b\<rangle>,r)) = cons(a, domain(r))"
 by blast
 
 lemma domain_Un_eq [simp]: "domain(A \<union> B) = domain(A) \<union> domain(B)"
@@ -620,12 +620,12 @@
 
 (** Range **)
 
-lemma rangeI [intro]: "<a,b>\<in> r \<Longrightarrow> b \<in> range(r)"
+lemma rangeI [intro]: "\<langle>a,b\<rangle>\<in> r \<Longrightarrow> b \<in> range(r)"
 apply (unfold range_def)
 apply (erule converseI [THEN domainI])
 done
 
-lemma rangeE [elim!]: "\<lbrakk>b \<in> range(r);  \<And>x. <x,b>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+lemma rangeE [elim!]: "\<lbrakk>b \<in> range(r);  \<And>x. \<langle>x,b\<rangle>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (unfold range_def, blast)
 
 lemma range_subset: "range(A*B) \<subseteq> B"
@@ -640,7 +640,7 @@
 lemma range_0 [simp]: "range(0) = 0"
 by blast
 
-lemma range_cons [simp]: "range(cons(<a,b>,r)) = cons(b, range(r))"
+lemma range_cons [simp]: "range(cons(\<langle>a,b\<rangle>,r)) = cons(b, range(r))"
 by blast
 
 lemma range_Un_eq [simp]: "range(A \<union> B) = range(A) \<union> range(B)"
@@ -661,21 +661,21 @@
 
 (** Field **)
 
-lemma fieldI1: "<a,b>\<in> r \<Longrightarrow> a \<in> field(r)"
+lemma fieldI1: "\<langle>a,b\<rangle>\<in> r \<Longrightarrow> a \<in> field(r)"
 by (unfold field_def, blast)
 
-lemma fieldI2: "<a,b>\<in> r \<Longrightarrow> b \<in> field(r)"
+lemma fieldI2: "\<langle>a,b\<rangle>\<in> r \<Longrightarrow> b \<in> field(r)"
 by (unfold field_def, blast)
 
 lemma fieldCI [intro]:
-    "(\<not> <c,a>\<in>r \<Longrightarrow> <a,b>\<in> r) \<Longrightarrow> a \<in> field(r)"
+    "(\<not> \<langle>c,a\<rangle>\<in>r \<Longrightarrow> \<langle>a,b\<rangle>\<in> r) \<Longrightarrow> a \<in> field(r)"
 apply (unfold field_def, blast)
 done
 
 lemma fieldE [elim!]:
      "\<lbrakk>a \<in> field(r);
-         \<And>x. <a,x>\<in> r \<Longrightarrow> P;
-         \<And>x. <x,a>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+         \<And>x. \<langle>a,x\<rangle>\<in> r \<Longrightarrow> P;
+         \<And>x. \<langle>x,a\<rangle>\<in> r \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (unfold field_def, blast)
 
 lemma field_subset: "field(A*B) \<subseteq> A \<union> B"
@@ -706,7 +706,7 @@
 lemma field_0 [simp]: "field(0) = 0"
 by blast
 
-lemma field_cons [simp]: "field(cons(<a,b>,r)) = cons(a, cons(b, field(r)))"
+lemma field_cons [simp]: "field(cons(\<langle>a,b\<rangle>,r)) = cons(a, cons(b, field(r)))"
 by blast
 
 lemma field_Un_eq [simp]: "field(A \<union> B) = field(A) \<union> field(B)"
@@ -730,26 +730,26 @@
 lemma rel_Un: "\<lbrakk>r \<subseteq> A*B;  s \<subseteq> C*D\<rbrakk> \<Longrightarrow> (r \<union> s) \<subseteq> (A \<union> C) * (B \<union> D)"
 by blast
 
-lemma domain_Diff_eq: "\<lbrakk><a,c> \<in> r; c\<noteq>b\<rbrakk> \<Longrightarrow> domain(r-{<a,b>}) = domain(r)"
+lemma domain_Diff_eq: "\<lbrakk>\<langle>a,c\<rangle> \<in> r; c\<noteq>b\<rbrakk> \<Longrightarrow> domain(r-{\<langle>a,b\<rangle>}) = domain(r)"
 by blast
 
-lemma range_Diff_eq: "\<lbrakk><c,b> \<in> r; c\<noteq>a\<rbrakk> \<Longrightarrow> range(r-{<a,b>}) = range(r)"
+lemma range_Diff_eq: "\<lbrakk>\<langle>c,b\<rangle> \<in> r; c\<noteq>a\<rbrakk> \<Longrightarrow> range(r-{\<langle>a,b\<rangle>}) = range(r)"
 by blast
 
 
 subsection\<open>Image of a Set under a Function or Relation\<close>
 
-lemma image_iff: "b \<in> r``A \<longleftrightarrow> (\<exists>x\<in>A. <x,b>\<in>r)"
+lemma image_iff: "b \<in> r``A \<longleftrightarrow> (\<exists>x\<in>A. \<langle>x,b\<rangle>\<in>r)"
 by (unfold image_def, blast)
 
-lemma image_singleton_iff: "b \<in> r``{a} \<longleftrightarrow> <a,b>\<in>r"
+lemma image_singleton_iff: "b \<in> r``{a} \<longleftrightarrow> \<langle>a,b\<rangle>\<in>r"
 by (rule image_iff [THEN iff_trans], blast)
 
-lemma imageI [intro]: "\<lbrakk><a,b>\<in> r;  a\<in>A\<rbrakk> \<Longrightarrow> b \<in> r``A"
+lemma imageI [intro]: "\<lbrakk>\<langle>a,b\<rangle>\<in> r;  a\<in>A\<rbrakk> \<Longrightarrow> b \<in> r``A"
 by (unfold image_def, blast)
 
 lemma imageE [elim!]:
-    "\<lbrakk>b: r``A;  \<And>x.\<lbrakk><x,b>\<in> r;  x\<in>A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+    "\<lbrakk>b: r``A;  \<And>x.\<lbrakk>\<langle>x,b\<rangle>\<in> r;  x\<in>A\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 by (unfold image_def, blast)
 
 lemma image_subset: "r \<subseteq> A*B \<Longrightarrow> r``C \<subseteq> B"
@@ -765,7 +765,7 @@
 by blast
 
 lemma Collect_image_eq:
-     "{z \<in> Sigma(A,B). P(z)} `` C = (\<Union>x \<in> A. {y \<in> B(x). x \<in> C \<and> P(<x,y>)})"
+     "{z \<in> Sigma(A,B). P(z)} `` C = (\<Union>x \<in> A. {y \<in> B(x). x \<in> C \<and> P(\<langle>x,y\<rangle>)})"
 by blast
 
 lemma image_Int_subset: "r``(A \<inter> B) \<subseteq> (r``A) \<inter> (r``B)"
@@ -792,17 +792,17 @@
 subsection\<open>Inverse Image of a Set under a Function or Relation\<close>
 
 lemma vimage_iff:
-    "a \<in> r-``B \<longleftrightarrow> (\<exists>y\<in>B. <a,y>\<in>r)"
+    "a \<in> r-``B \<longleftrightarrow> (\<exists>y\<in>B. \<langle>a,y\<rangle>\<in>r)"
 by (unfold vimage_def image_def converse_def, blast)
 
-lemma vimage_singleton_iff: "a \<in> r-``{b} \<longleftrightarrow> <a,b>\<in>r"
+lemma vimage_singleton_iff: "a \<in> r-``{b} \<longleftrightarrow> \<langle>a,b\<rangle>\<in>r"
 by (rule vimage_iff [THEN iff_trans], blast)
 
-lemma vimageI [intro]: "\<lbrakk><a,b>\<in> r;  b\<in>B\<rbrakk> \<Longrightarrow> a \<in> r-``B"
+lemma vimageI [intro]: "\<lbrakk>\<langle>a,b\<rangle>\<in> r;  b\<in>B\<rbrakk> \<Longrightarrow> a \<in> r-``B"
 by (unfold vimage_def, blast)
 
 lemma vimageE [elim!]:
-    "\<lbrakk>a: r-``B;  \<And>x.\<lbrakk><a,x>\<in> r;  x\<in>B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
+    "\<lbrakk>a: r-``B;  \<And>x.\<lbrakk>\<langle>a,x\<rangle>\<in> r;  x\<in>B\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
 apply (unfold vimage_def, blast)
 done
 
@@ -941,11 +941,11 @@
 by blast
 
 lemma Collect_Collect_eq [simp]:
-     "Collect(Collect(A,P), Q) = Collect(A, %x. P(x) \<and> Q(x))"
+     "Collect(Collect(A,P), Q) = Collect(A, \<lambda>x. P(x) \<and> Q(x))"
 by blast
 
 lemma Collect_Int_Collect_eq:
-     "Collect(A,P) \<inter> Collect(A,Q) = Collect(A, %x. P(x) \<and> Q(x))"
+     "Collect(A,P) \<inter> Collect(A,Q) = Collect(A, \<lambda>x. P(x) \<and> Q(x))"
 by blast
 
 lemma Collect_Union_eq [simp]:
--- a/src/ZF/ex/CoUnit.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/CoUnit.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -59,13 +59,13 @@
   \<comment> \<open>Proving freeness results.\<close>
   by (fast elim!: counit2.free_elims)
 
-lemma Con2_bnd_mono: "bnd_mono(univ(0), %x. Con2(x, x))"
+lemma Con2_bnd_mono: "bnd_mono(univ(0), \<lambda>x. Con2(x, x))"
   apply (unfold counit2.con_defs)
   apply (rule bnd_monoI)
    apply (assumption | rule subset_refl QPair_subset_univ QPair_mono)+
   done
 
-lemma lfp_Con2_in_counit2: "lfp(univ(0), %x. Con2(x,x)) \<in> counit2"
+lemma lfp_Con2_in_counit2: "lfp(univ(0), \<lambda>x. Con2(x,x)) \<in> counit2"
   apply (rule singletonI [THEN counit2.coinduct])
   apply (rule qunivI [THEN singleton_subsetI])
   apply (rule subset_trans [OF lfp_subset empty_subsetI [THEN univ_mono]])
@@ -90,7 +90,7 @@
   apply (assumption | rule conjI counit2_Int_Vset_subset [THEN Int_Vset_subset])+
   done
 
-lemma counit2_eq_univ: "counit2 = {lfp(univ(0), %x. Con2(x,x))}"
+lemma counit2_eq_univ: "counit2 = {lfp(univ(0), \<lambda>x. Con2(x,x))}"
   apply (rule equalityI)
    apply (rule_tac [2] lfp_Con2_in_counit2 [THEN singleton_subsetI])
   apply (rule subsetI)
--- a/src/ZF/ex/Commutation.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/Commutation.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,29 +8,29 @@
 theory Commutation imports ZF begin
 
 definition
-  square  :: "[i, i, i, i] => o" where
+  square  :: "[i, i, i, i] \<Rightarrow> o" where
   "square(r,s,t,u) \<equiv>
-    (\<forall>a b. <a,b> \<in> r \<longrightarrow> (\<forall>c. <a, c> \<in> s \<longrightarrow> (\<exists>x. <b,x> \<in> t \<and> <c,x> \<in> u)))"
+    (\<forall>a b. \<langle>a,b\<rangle> \<in> r \<longrightarrow> (\<forall>c. \<langle>a, c\<rangle> \<in> s \<longrightarrow> (\<exists>x. \<langle>b,x\<rangle> \<in> t \<and> \<langle>c,x\<rangle> \<in> u)))"
 
 definition
-  commute :: "[i, i] => o" where
+  commute :: "[i, i] \<Rightarrow> o" where
   "commute(r,s) \<equiv> square(r,s,s,r)"
 
 definition
-  diamond :: "i=>o" where
+  diamond :: "i\<Rightarrow>o" where
   "diamond(r)   \<equiv> commute(r, r)"
 
 definition
-  strip :: "i=>o" where
+  strip :: "i\<Rightarrow>o" where
   "strip(r) \<equiv> commute(r^*, r)"
 
 definition
-  Church_Rosser :: "i => o" where
-  "Church_Rosser(r) \<equiv> (\<forall>x y. <x,y> \<in>  (r \<union> converse(r))^* \<longrightarrow>
-                        (\<exists>z. <x,z> \<in> r^* \<and> <y,z> \<in> r^*))"
+  Church_Rosser :: "i \<Rightarrow> o" where
+  "Church_Rosser(r) \<equiv> (\<forall>x y. \<langle>x,y\<rangle> \<in>  (r \<union> converse(r))^* \<longrightarrow>
+                        (\<exists>z. \<langle>x,z\<rangle> \<in> r^* \<and> \<langle>y,z\<rangle> \<in> r^*))"
 
 definition
-  confluent :: "i=>o" where
+  confluent :: "i\<Rightarrow>o" where
   "confluent(r) \<equiv> diamond(r^*)"
 
 
--- a/src/ZF/ex/Group.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/Group.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -13,28 +13,28 @@
 (*First, we must simulate a record declaration:
 record monoid =
   carrier :: i
-  mult :: "[i,i] => i" (infixl "\<cdot>\<index>" 70)
+  mult :: "[i,i] \<Rightarrow> i" (infixl "\<cdot>\<index>" 70)
   one :: i ("\<one>\<index>")
 *)
 
 definition
-  carrier :: "i => i" where
+  carrier :: "i \<Rightarrow> i" where
   "carrier(M) \<equiv> fst(M)"
 
 definition
-  mmult :: "[i, i, i] => i" (infixl \<open>\<cdot>\<index>\<close> 70) where
-  "mmult(M,x,y) \<equiv> fst(snd(M)) ` <x,y>"
+  mmult :: "[i, i, i] \<Rightarrow> i" (infixl \<open>\<cdot>\<index>\<close> 70) where
+  "mmult(M,x,y) \<equiv> fst(snd(M)) ` \<langle>x,y\<rangle>"
 
 definition
-  one :: "i => i" (\<open>\<one>\<index>\<close>) where
+  one :: "i \<Rightarrow> i" (\<open>\<one>\<index>\<close>) where
   "one(M) \<equiv> fst(snd(snd(M)))"
 
 definition
-  update_carrier :: "[i,i] => i" where
+  update_carrier :: "[i,i] \<Rightarrow> i" where
   "update_carrier(M,A) \<equiv> <A,snd(M)>"
 
 definition
-  m_inv :: "i => i => i" (\<open>inv\<index> _\<close> [81] 80) where
+  m_inv :: "i \<Rightarrow> i \<Rightarrow> i" (\<open>inv\<index> _\<close> [81] 80) where
   "inv\<^bsub>G\<^esub> x \<equiv> (THE y. y \<in> carrier(G) \<and> y \<cdot>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub> \<and> x \<cdot>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub>)"
 
 locale monoid = fixes G (structure)
@@ -48,16 +48,16 @@
       and r_one [simp]: "x \<in> carrier(G) \<Longrightarrow> x \<cdot> \<one> = x"
 
 text\<open>Simulating the record\<close>
-lemma carrier_eq [simp]: "carrier(<A,Z>) = A"
+lemma carrier_eq [simp]: "carrier(\<langle>A,Z\<rangle>) = A"
   by (simp add: carrier_def)
 
-lemma mult_eq [simp]: "mmult(<A,M,Z>, x, y) = M ` <x,y>"
+lemma mult_eq [simp]: "mmult(<A,M,Z>, x, y) = M ` \<langle>x,y\<rangle>"
   by (simp add: mmult_def)
 
 lemma one_eq [simp]: "one(<A,M,I,Z>) = I"
   by (simp add: one_def)
 
-lemma update_carrier_eq [simp]: "update_carrier(<A,Z>,B) = <B,Z>"
+lemma update_carrier_eq [simp]: "update_carrier(\<langle>A,Z\<rangle>,B) = \<langle>B,Z\<rangle>"
   by (simp add: update_carrier_def)
 
 lemma carrier_update_carrier [simp]: "carrier(update_carrier(M,B)) = B"
@@ -302,9 +302,9 @@
 subsection \<open>Direct Products\<close>
 
 definition
-  DirProdGroup :: "[i,i] => i"  (infixr \<open>\<Otimes>\<close> 80) where
+  DirProdGroup :: "[i,i] \<Rightarrow> i"  (infixr \<open>\<Otimes>\<close> 80) where
   "G \<Otimes> H \<equiv> <carrier(G) \<times> carrier(H),
-              (\<lambda><<g,h>, <g', h'>>
+              (\<lambda><\<langle>g,h\<rangle>, <g', h'>>
                    \<in> (carrier(G) \<times> carrier(H)) \<times> (carrier(G) \<times> carrier(H)).
                 <g \<cdot>\<^bsub>G\<^esub> g', h \<cdot>\<^bsub>H\<^esub> h'>),
               <\<one>\<^bsub>G\<^esub>, \<one>\<^bsub>H\<^esub>>, 0>"
@@ -329,14 +329,14 @@
 
 lemma mult_DirProdGroup [simp]:
      "\<lbrakk>g \<in> carrier(G); h \<in> carrier(H); g' \<in> carrier(G); h' \<in> carrier(H)\<rbrakk>
-      \<Longrightarrow> <g, h> \<cdot>\<^bsub>G \<Otimes> H\<^esub> <g', h'> = <g \<cdot>\<^bsub>G\<^esub> g', h \<cdot>\<^bsub>H\<^esub> h'>"
+      \<Longrightarrow> \<langle>g, h\<rangle> \<cdot>\<^bsub>G \<Otimes> H\<^esub> <g', h'> = <g \<cdot>\<^bsub>G\<^esub> g', h \<cdot>\<^bsub>H\<^esub> h'>"
   by (simp add: DirProdGroup_def)
 
 lemma inv_DirProdGroup [simp]:
   assumes "group(G)" and "group(H)"
   assumes g: "g \<in> carrier(G)"
       and h: "h \<in> carrier(H)"
-  shows "inv \<^bsub>G \<Otimes> H\<^esub> <g, h> = <inv\<^bsub>G\<^esub> g, inv\<^bsub>H\<^esub> h>"
+  shows "inv \<^bsub>G \<Otimes> H\<^esub> \<langle>g, h\<rangle> = <inv\<^bsub>G\<^esub> g, inv\<^bsub>H\<^esub> h>"
   apply (rule group.inv_equality [OF DirProdGroup_group])
   apply (simp_all add: assms group.l_inv)
   done
@@ -344,7 +344,7 @@
 subsection \<open>Isomorphisms\<close>
 
 definition
-  hom :: "[i,i] => i" where
+  hom :: "[i,i] \<Rightarrow> i" where
   "hom(G,H) \<equiv>
     {h \<in> carrier(G) -> carrier(H).
       (\<forall>x \<in> carrier(G). \<forall>y \<in> carrier(G). h ` (x \<cdot>\<^bsub>G\<^esub> y) = (h ` x) \<cdot>\<^bsub>H\<^esub> (h ` y))}"
@@ -370,7 +370,7 @@
 subsection \<open>Isomorphisms\<close>
 
 definition
-  iso :: "[i,i] => i"  (infixr \<open>\<cong>\<close> 60) where
+  iso :: "[i,i] \<Rightarrow> i"  (infixr \<open>\<cong>\<close> 60) where
   "G \<cong> H \<equiv> hom(G,H) \<inter> bij(carrier(G), carrier(H))"
 
 lemma (in group) iso_refl: "id(carrier(G)) \<in> G \<cong> G"
@@ -392,7 +392,7 @@
 
 lemma DirProdGroup_commute_iso:
   assumes "group(G)" and "group(H)"
-  shows "(\<lambda><x,y> \<in> carrier(G \<Otimes> H). <y,x>) \<in> (G \<Otimes> H) \<cong> (H \<Otimes> G)"
+  shows "(\<lambda>\<langle>x,y\<rangle> \<in> carrier(G \<Otimes> H). \<langle>y,x\<rangle>) \<in> (G \<Otimes> H) \<cong> (H \<Otimes> G)"
 proof -
   interpret group G by fact
   interpret group H by fact
@@ -401,7 +401,7 @@
 
 lemma DirProdGroup_assoc_iso:
   assumes "group(G)" and "group(H)" and "group(I)"
-  shows "(\<lambda><<x,y>,z> \<in> carrier((G \<Otimes> H) \<Otimes> I). <x,<y,z>>)
+  shows "(\<lambda><\<langle>x,y\<rangle>,z> \<in> carrier((G \<Otimes> H) \<Otimes> I). <x,\<langle>y,z\<rangle>>)
           \<in> ((G \<Otimes> H) \<Otimes> I) \<cong> (G \<Otimes> (H \<Otimes> I))"
 proof -
   interpret group G by fact
@@ -502,10 +502,10 @@
 subsection \<open>Bijections of a Set, Permutation Groups, Automorphism Groups\<close>
 
 definition
-  BijGroup :: "i=>i" where
+  BijGroup :: "i\<Rightarrow>i" where
   "BijGroup(S) \<equiv>
     <bij(S,S),
-     \<lambda><g,f> \<in> bij(S,S) \<times> bij(S,S). g O f,
+     \<lambda>\<langle>g,f\<rangle> \<in> bij(S,S) \<times> bij(S,S). g O f,
      id(S), 0>"
 
 
@@ -537,11 +537,11 @@
 
 
 definition
-  auto :: "i=>i" where
+  auto :: "i\<Rightarrow>i" where
   "auto(G) \<equiv> iso(G,G)"
 
 definition
-  AutoGroup :: "i=>i" where
+  AutoGroup :: "i\<Rightarrow>i" where
   "AutoGroup(G) \<equiv> update_carrier(BijGroup(carrier(G)), auto(G))"
 
 
@@ -576,23 +576,23 @@
 subsection\<open>Cosets and Quotient Groups\<close>
 
 definition
-  r_coset  :: "[i,i,i] => i"  (infixl \<open>#>\<index>\<close> 60) where
+  r_coset  :: "[i,i,i] \<Rightarrow> i"  (infixl \<open>#>\<index>\<close> 60) where
   "H #>\<^bsub>G\<^esub> a \<equiv> \<Union>h\<in>H. {h \<cdot>\<^bsub>G\<^esub> a}"
 
 definition
-  l_coset  :: "[i,i,i] => i"  (infixl \<open><#\<index>\<close> 60) where
+  l_coset  :: "[i,i,i] \<Rightarrow> i"  (infixl \<open><#\<index>\<close> 60) where
   "a <#\<^bsub>G\<^esub> H \<equiv> \<Union>h\<in>H. {a \<cdot>\<^bsub>G\<^esub> h}"
 
 definition
-  RCOSETS  :: "[i,i] => i"  (\<open>rcosets\<index> _\<close> [81] 80) where
+  RCOSETS  :: "[i,i] \<Rightarrow> i"  (\<open>rcosets\<index> _\<close> [81] 80) where
   "rcosets\<^bsub>G\<^esub> H \<equiv> \<Union>a\<in>carrier(G). {H #>\<^bsub>G\<^esub> a}"
 
 definition
-  set_mult :: "[i,i,i] => i"  (infixl \<open><#>\<index>\<close> 60) where
+  set_mult :: "[i,i,i] \<Rightarrow> i"  (infixl \<open><#>\<index>\<close> 60) where
   "H <#>\<^bsub>G\<^esub> K \<equiv> \<Union>h\<in>H. \<Union>k\<in>K. {h \<cdot>\<^bsub>G\<^esub> k}"
 
 definition
-  SET_INV  :: "[i,i] => i"  (\<open>set'_inv\<index> _\<close> [81] 80) where
+  SET_INV  :: "[i,i] \<Rightarrow> i"  (\<open>set'_inv\<index> _\<close> [81] 80) where
   "set_inv\<^bsub>G\<^esub> H \<equiv> \<Union>h\<in>H. {inv\<^bsub>G\<^esub> h}"
 
 
@@ -860,8 +860,8 @@
 subsubsection\<open>Two distinct right cosets are disjoint\<close>
 
 definition
-  r_congruent :: "[i,i] => i" (\<open>rcong\<index> _\<close> [60] 60) where
-  "rcong\<^bsub>G\<^esub> H \<equiv> {<x,y> \<in> carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \<cdot>\<^bsub>G\<^esub> y \<in> H}"
+  r_congruent :: "[i,i] \<Rightarrow> i" (\<open>rcong\<index> _\<close> [60] 60) where
+  "rcong\<^bsub>G\<^esub> H \<equiv> {\<langle>x,y\<rangle> \<in> carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \<cdot>\<^bsub>G\<^esub> y \<in> H}"
 
 
 lemma (in subgroup) equiv_rcong:
@@ -940,7 +940,7 @@
 subsection \<open>Order of a Group and Lagrange's Theorem\<close>
 
 definition
-  order :: "i => i" where
+  order :: "i \<Rightarrow> i" where
   "order(S) \<equiv> |carrier(S)|"
 
 lemma (in group) rcos_self:
@@ -1020,10 +1020,10 @@
 subsection \<open>Quotient Groups: Factorization of a Group\<close>
 
 definition
-  FactGroup :: "[i,i] => i" (infixl \<open>Mod\<close> 65) where
+  FactGroup :: "[i,i] \<Rightarrow> i" (infixl \<open>Mod\<close> 65) where
     \<comment> \<open>Actually defined for groups rather than monoids\<close>
   "G Mod H \<equiv>
-     <rcosets\<^bsub>G\<^esub> H, \<lambda><K1,K2> \<in> (rcosets\<^bsub>G\<^esub> H) \<times> (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>"
+     <rcosets\<^bsub>G\<^esub> H, \<lambda>\<langle>K1,K2\<rangle> \<in> (rcosets\<^bsub>G\<^esub> H) \<times> (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>"
 
 lemma (in normal) setmult_closed:
      "\<lbrakk>K1 \<in> rcosets H; K2 \<in> rcosets H\<rbrakk> \<Longrightarrow> K1 <#> K2 \<in> rcosets H"
@@ -1084,7 +1084,7 @@
   range of that homomorphism.\<close>
 
 definition
-  kernel :: "[i,i,i] => i" where
+  kernel :: "[i,i,i] \<Rightarrow> i" where
     \<comment> \<open>the kernel of a homomorphism\<close>
   "kernel(G,H,h) \<equiv> {x \<in> carrier(G). h ` x = \<one>\<^bsub>H\<^esub>}"
 
--- a/src/ZF/ex/LList.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/LList.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -16,7 +16,7 @@
 theory LList imports ZF begin
 
 consts
-  llist  :: "i=>i"
+  llist  :: "i\<Rightarrow>i"
 
 codatatype
   "llist(A)" = LNil | LCons ("a \<in> A", "l \<in> llist(A)")
@@ -24,7 +24,7 @@
 
 (*Coinductive definition of equality*)
 consts
-  lleq :: "i=>i"
+  lleq :: "i\<Rightarrow>i"
 
 (*Previously used <*> in the domain and variant pairs as elements.  But
   standard pairs work just as well.  To use variant pairs, must change prefix
@@ -32,7 +32,7 @@
 coinductive
   domains "lleq(A)" \<subseteq> "llist(A) * llist(A)"
   intros
-    LNil:  "<LNil, LNil> \<in> lleq(A)"
+    LNil:  "\<langle>LNil, LNil\<rangle> \<in> lleq(A)"
     LCons: "\<lbrakk>a \<in> A; <l,l'> \<in> lleq(A)\<rbrakk> 
             \<Longrightarrow> <LCons(a,l), LCons(a,l')> \<in> lleq(A)"
   type_intros  llist.intros
@@ -40,10 +40,10 @@
 
 (*Lazy list functions; flip is not definitional!*)
 definition
-  lconst   :: "i => i"  where
-  "lconst(a) \<equiv> lfp(univ(a), %l. LCons(a,l))"
+  lconst   :: "i \<Rightarrow> i"  where
+  "lconst(a) \<equiv> lfp(univ(a), \<lambda>l. LCons(a,l))"
 
-axiomatization flip :: "i => i"
+axiomatization flip :: "i \<Rightarrow> i"
 where
   flip_LNil:   "flip(LNil) = LNil" and
   flip_LCons:  "\<lbrakk>x \<in> bool; l \<in> llist(bool)\<rbrakk> 
@@ -170,7 +170,7 @@
 
 lemma equal_llist_implies_leq:
      "\<lbrakk>l=l';  l \<in> llist(A)\<rbrakk> \<Longrightarrow> <l,l'> \<in> lleq(A)"
-apply (rule_tac X = "{<l,l>. l \<in> llist (A) }" in lleq.coinduct)
+apply (rule_tac X = "{\<langle>l,l\<rangle>. l \<in> llist (A) }" in lleq.coinduct)
 apply blast
 apply safe
 apply (erule_tac a=la in llist.cases, fast+)
@@ -183,7 +183,7 @@
 
 (*** lconst -- defined directly using lfp, but equivalent to a LList_corec ***)
 
-lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), %l. LCons(a,l))"
+lemma lconst_fun_bnd_mono: "bnd_mono(univ(a), \<lambda>l. LCons(a,l))"
 apply (unfold llist.con_defs )
 apply (rule bnd_monoI)
 apply (blast intro: A_subset_univ QInr_subset_univ)
--- a/src/ZF/ex/Limit.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/Limit.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -20,15 +20,15 @@
 theory Limit  imports  ZF begin
 
 definition
-  rel :: "[i,i,i]=>o"  where
-    "rel(D,x,y) \<equiv> <x,y>:snd(D)"
+  rel :: "[i,i,i]\<Rightarrow>o"  where
+    "rel(D,x,y) \<equiv> \<langle>x,y\<rangle>:snd(D)"
 
 definition
-  set :: "i=>i"  where
+  set :: "i\<Rightarrow>i"  where
     "set(D) \<equiv> fst(D)"
 
 definition
-  po  :: "i=>o"  where
+  po  :: "i\<Rightarrow>o"  where
     "po(D) \<equiv>
      (\<forall>x \<in> set(D). rel(D,x,x)) \<and>
      (\<forall>x \<in> set(D). \<forall>y \<in> set(D). \<forall>z \<in> set(D).
@@ -36,66 +36,66 @@
      (\<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) \<longrightarrow> rel(D,y,x) \<longrightarrow> x = y)"
 
 definition
-  chain :: "[i,i]=>o"  where
+  chain :: "[i,i]\<Rightarrow>o"  where
     (* Chains are object level functions nat->set(D) *)
     "chain(D,X) \<equiv> X \<in> nat->set(D) \<and> (\<forall>n \<in> nat. rel(D,X`n,X`(succ(n))))"
 
 definition
-  isub :: "[i,i,i]=>o"  where
+  isub :: "[i,i,i]\<Rightarrow>o"  where
     "isub(D,X,x) \<equiv> x \<in> set(D) \<and> (\<forall>n \<in> nat. rel(D,X`n,x))"
 
 definition
-  islub :: "[i,i,i]=>o"  where
+  islub :: "[i,i,i]\<Rightarrow>o"  where
     "islub(D,X,x) \<equiv> isub(D,X,x) \<and> (\<forall>y. isub(D,X,y) \<longrightarrow> rel(D,x,y))"
 
 definition
-  lub :: "[i,i]=>i"  where
+  lub :: "[i,i]\<Rightarrow>i"  where
     "lub(D,X) \<equiv> THE x. islub(D,X,x)"
 
 definition
-  cpo :: "i=>o"  where
+  cpo :: "i\<Rightarrow>o"  where
     "cpo(D) \<equiv> po(D) \<and> (\<forall>X. chain(D,X) \<longrightarrow> (\<exists>x. islub(D,X,x)))"
 
 definition
-  pcpo :: "i=>o"  where
+  pcpo :: "i\<Rightarrow>o"  where
     "pcpo(D) \<equiv> cpo(D) \<and> (\<exists>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y))"
 
 definition
-  bot :: "i=>i"  where
+  bot :: "i\<Rightarrow>i"  where
     "bot(D) \<equiv> THE x. x \<in> set(D) \<and> (\<forall>y \<in> set(D). rel(D,x,y))"
 
 definition
-  mono :: "[i,i]=>i"  where
+  mono :: "[i,i]\<Rightarrow>i"  where
     "mono(D,E) \<equiv>
      {f \<in> set(D)->set(E).
       \<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) \<longrightarrow> rel(E,f`x,f`y)}"
 
 definition
-  cont :: "[i,i]=>i"  where
+  cont :: "[i,i]\<Rightarrow>i"  where
     "cont(D,E) \<equiv>
      {f \<in> mono(D,E).
       \<forall>X. chain(D,X) \<longrightarrow> f`(lub(D,X)) = lub(E,\<lambda>n \<in> nat. f`(X`n))}"
 
 definition
-  cf :: "[i,i]=>i"  where
+  cf :: "[i,i]\<Rightarrow>i"  where
     "cf(D,E) \<equiv>
      <cont(D,E),
       {y \<in> cont(D,E)*cont(D,E). \<forall>x \<in> set(D). rel(E,(fst(y))`x,(snd(y))`x)}>"
 
 definition
-  suffix :: "[i,i]=>i"  where
+  suffix :: "[i,i]\<Rightarrow>i"  where
     "suffix(X,n) \<equiv> \<lambda>m \<in> nat. X`(n #+ m)"
 
 definition
-  subchain :: "[i,i]=>o"  where
+  subchain :: "[i,i]\<Rightarrow>o"  where
     "subchain(X,Y) \<equiv> \<forall>m \<in> nat. \<exists>n \<in> nat. X`m = Y`(m #+ n)"
 
 definition
-  dominate :: "[i,i,i]=>o"  where
+  dominate :: "[i,i,i]\<Rightarrow>o"  where
     "dominate(D,X,Y) \<equiv> \<forall>m \<in> nat. \<exists>n \<in> nat. rel(D,X`m,Y`n)"
 
 definition
-  matrix :: "[i,i]=>o"  where
+  matrix :: "[i,i]\<Rightarrow>o"  where
     "matrix(D,M) \<equiv>
      M \<in> nat -> (nat -> set(D)) \<and>
      (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`succ(n)`m)) \<and>
@@ -103,89 +103,89 @@
      (\<forall>n \<in> nat. \<forall>m \<in> nat. rel(D,M`n`m,M`succ(n)`succ(m)))"
 
 definition
-  projpair  :: "[i,i,i,i]=>o"  where
+  projpair  :: "[i,i,i,i]\<Rightarrow>o"  where
     "projpair(D,E,e,p) \<equiv>
      e \<in> cont(D,E) \<and> p \<in> cont(E,D) \<and>
      p O e = id(set(D)) \<and> rel(cf(E,E),e O p,id(set(E)))"
 
 definition
-  emb       :: "[i,i,i]=>o"  where
+  emb       :: "[i,i,i]\<Rightarrow>o"  where
     "emb(D,E,e) \<equiv> \<exists>p. projpair(D,E,e,p)"
 
 definition
-  Rp        :: "[i,i,i]=>i"  where
+  Rp        :: "[i,i,i]\<Rightarrow>i"  where
     "Rp(D,E,e) \<equiv> THE p. projpair(D,E,e,p)"
 
 definition
   (* Twice, constructions on cpos are more difficult. *)
-  iprod     :: "i=>i"  where
+  iprod     :: "i\<Rightarrow>i"  where
     "iprod(DD) \<equiv>
      <(\<Prod>n \<in> nat. set(DD`n)),
       {x:(\<Prod>n \<in> nat. set(DD`n))*(\<Prod>n \<in> nat. set(DD`n)).
        \<forall>n \<in> nat. rel(DD`n,fst(x)`n,snd(x)`n)}>"
 
 definition
-  mkcpo     :: "[i,i=>o]=>i"  where
+  mkcpo     :: "[i,i\<Rightarrow>o]\<Rightarrow>i"  where
     (* Cannot use rel(D), is meta fun, need two more args *)
     "mkcpo(D,P) \<equiv>
      <{x \<in> set(D). P(x)},{x \<in> set(D)*set(D). rel(D,fst(x),snd(x))}>"
 
 definition
-  subcpo    :: "[i,i]=>o"  where
+  subcpo    :: "[i,i]\<Rightarrow>o"  where
     "subcpo(D,E) \<equiv>
      set(D) \<subseteq> set(E) \<and>
      (\<forall>x \<in> set(D). \<forall>y \<in> set(D). rel(D,x,y) \<longleftrightarrow> rel(E,x,y)) \<and>
      (\<forall>X. chain(D,X) \<longrightarrow> lub(E,X):set(D))"
 
 definition
-  subpcpo   :: "[i,i]=>o"  where
+  subpcpo   :: "[i,i]\<Rightarrow>o"  where
     "subpcpo(D,E) \<equiv> subcpo(D,E) \<and> bot(E):set(D)"
 
 definition
-  emb_chain :: "[i,i]=>o"  where
+  emb_chain :: "[i,i]\<Rightarrow>o"  where
     "emb_chain(DD,ee) \<equiv>
      (\<forall>n \<in> nat. cpo(DD`n)) \<and> (\<forall>n \<in> nat. emb(DD`n,DD`succ(n),ee`n))"
 
 definition
-  Dinf      :: "[i,i]=>i"  where
+  Dinf      :: "[i,i]\<Rightarrow>i"  where
     "Dinf(DD,ee) \<equiv>
      mkcpo(iprod(DD))
-     (%x. \<forall>n \<in> nat. Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n)"
+     (\<lambda>x. \<forall>n \<in> nat. Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n)"
 
 definition
-  e_less    :: "[i,i,i,i]=>i"  where
+  e_less    :: "[i,i,i,i]\<Rightarrow>i"  where
   (* Valid for m \<le> n only. *)
-    "e_less(DD,ee,m,n) \<equiv> rec(n#-m,id(set(DD`m)),%x y. ee`(m#+x) O y)"
+    "e_less(DD,ee,m,n) \<equiv> rec(n#-m,id(set(DD`m)),\<lambda>x y. ee`(m#+x) O y)"
 
 
 definition
-  e_gr      :: "[i,i,i,i]=>i"  where
+  e_gr      :: "[i,i,i,i]\<Rightarrow>i"  where
   (* Valid for n \<le> m only. *)
     "e_gr(DD,ee,m,n) \<equiv>
      rec(m#-n,id(set(DD`n)),
-         %x y. y O Rp(DD`(n#+x),DD`(succ(n#+x)),ee`(n#+x)))"
+         \<lambda>x y. y O Rp(DD`(n#+x),DD`(succ(n#+x)),ee`(n#+x)))"
 
 
 definition
-  eps       :: "[i,i,i,i]=>i"  where
+  eps       :: "[i,i,i,i]\<Rightarrow>i"  where
     "eps(DD,ee,m,n) \<equiv> if(m \<le> n,e_less(DD,ee,m,n),e_gr(DD,ee,m,n))"
 
 definition
-  rho_emb   :: "[i,i,i]=>i"  where
+  rho_emb   :: "[i,i,i]\<Rightarrow>i"  where
     "rho_emb(DD,ee,n) \<equiv> \<lambda>x \<in> set(DD`n). \<lambda>m \<in> nat. eps(DD,ee,n,m)`x"
 
 definition
-  rho_proj  :: "[i,i,i]=>i"  where
+  rho_proj  :: "[i,i,i]\<Rightarrow>i"  where
     "rho_proj(DD,ee,n) \<equiv> \<lambda>x \<in> set(Dinf(DD,ee)). x`n"
 
 definition
-  commute   :: "[i,i,i,i=>i]=>o"  where
+  commute   :: "[i,i,i,i\<Rightarrow>i]\<Rightarrow>o"  where
     "commute(DD,ee,E,r) \<equiv>
      (\<forall>n \<in> nat. emb(DD`n,E,r(n))) \<and>
      (\<forall>m \<in> nat. \<forall>n \<in> nat. m \<le> n \<longrightarrow> r(n) O eps(DD,ee,m,n) = r(m))"
 
 definition
-  mediating :: "[i,i,i=>i,i=>i,i]=>o"  where
+  mediating :: "[i,i,i\<Rightarrow>i,i\<Rightarrow>i,i]\<Rightarrow>o"  where
     "mediating(E,G,r,f,t) \<equiv> emb(E,G,t) \<and> (\<forall>n \<in> nat. f(n) = t O r(n))"
 
 
@@ -199,10 +199,10 @@
 lemma set_I: "x \<in> fst(D) \<Longrightarrow> x \<in> set(D)"
 by (simp add: set_def)
 
-lemma rel_I: "<x,y>:snd(D) \<Longrightarrow> rel(D,x,y)"
+lemma rel_I: "\<langle>x,y\<rangle>:snd(D) \<Longrightarrow> rel(D,x,y)"
 by (simp add: rel_def)
 
-lemma rel_E: "rel(D,x,y) \<Longrightarrow> <x,y>:snd(D)"
+lemma rel_E: "rel(D,x,y) \<Longrightarrow> \<langle>x,y\<rangle>:snd(D)"
 by (simp add: rel_def)
 
 (*----------------------------------------------------------------------*)
@@ -936,7 +936,7 @@
 apply (subst comp_assoc)
 apply (blast intro:  cpo_cf cont_cf comp_mono comp_pres_cont
              dest: projpair_ep_cont)
-apply (rule_tac P = "%x. rel (cf (E,D),p O e' O p',x)"
+apply (rule_tac P = "\<lambda>x. rel (cf (E,D),p O e' O p',x)"
          in projpair_p_cont [THEN cont_fun, THEN comp_id, THEN subst],
        assumption)
 apply (rule comp_mono)
@@ -959,7 +959,7 @@
 apply (subst comp_assoc)
 apply (blast intro:  cpo_cf cont_cf comp_mono comp_pres_cont
              dest: projpair_ep_cont)
-apply (rule_tac P = "%x. rel (cf (D,E), (e O p) O e',x)"
+apply (rule_tac P = "\<lambda>x. rel (cf (D,E), (e O p) O e',x)"
          in projpair_e_cont [THEN cont_fun, THEN id_comp, THEN subst],
        assumption)
 apply (blast intro: cpo_cf cont_cf comp_pres_cont projpair_rel comp_mono
@@ -1108,8 +1108,8 @@
 apply (blast intro: lam_type chain_iprod [THEN cpo_lub, THEN islub_in])
 apply (rule rel_iprodI, simp)
 (*looks like something should be inserted into the assumptions!*)
-apply (rule_tac P = "%t. rel (DD`na,t,lub (DD`na,\<lambda>x \<in> nat. X`x`na))"
-            and b1 = "%n. X`n`na" in beta [THEN subst])
+apply (rule_tac P = "\<lambda>t. rel (DD`na,t,lub (DD`na,\<lambda>x \<in> nat. X`x`na))"
+            and b1 = "\<lambda>n. X`n`na" in beta [THEN subst])
 apply (simp del: beta_if
             add: chain_iprod [THEN cpo_lub, THEN islub_ub] iprodE
                 chain_in)+
@@ -1832,12 +1832,12 @@
 apply (rule_tac [3] comp_fun_apply [THEN subst])
 apply (rename_tac [5] y)
 apply (rule_tac [5] P =
-         "%z. rel(DD`succ(y),
+         "\<lambda>z. rel(DD`succ(y),
                   (ee`y O Rp(DD'(y)`y,DD'(y)`succ(y),ee'(y)`y)) ` (x`succ(y)),
                   z)" for DD' ee'
        in id_conv [THEN subst])
 apply (rule_tac [6] rel_cf)
-(* Dinf and cont_fun doesn't go well together, both Pi(_,%x._). *)
+(* Dinf and cont_fun doesn't go well together, both Pi(_,\<lambda>x._). *)
 (* solves 10 of 11 subgoals *)
 apply (assumption |
        rule Dinf_prod [THEN apply_type] cont_fun Rp_cont e_less_cont
--- a/src/ZF/ex/NatSum.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/NatSum.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -18,7 +18,7 @@
 
 theory NatSum imports ZF begin
 
-consts sum :: "[i=>i, i] => i"
+consts sum :: "[i\<Rightarrow>i, i] \<Rightarrow> i"
 primrec 
   "sum (f,0) = #0"
   "sum (f, succ(n)) = f($#n) $+ sum(f,n)"
@@ -27,41 +27,41 @@
 declare zdiff_zmult_distrib [simp] zdiff_zmult_distrib2 [simp]
 
 (*The sum of the first n odd numbers equals n squared.*)
-lemma sum_of_odds: "n \<in> nat \<Longrightarrow> sum (%i. i $+ i $+ #1, n) = $#n $* $#n"
+lemma sum_of_odds: "n \<in> nat \<Longrightarrow> sum (\<lambda>i. i $+ i $+ #1, n) = $#n $* $#n"
 by (induct_tac "n", auto)
 
 (*The sum of the first n odd squares*)
 lemma sum_of_odd_squares:
-     "n \<in> nat \<Longrightarrow> #3 $* sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) =  
+     "n \<in> nat \<Longrightarrow> #3 $* sum (\<lambda>i. (i $+ i $+ #1) $* (i $+ i $+ #1), n) =  
       $#n $* (#4 $* $#n $* $#n $- #1)"
 by (induct_tac "n", auto)
 
 (*The sum of the first n odd cubes*)
 lemma sum_of_odd_cubes:
      "n \<in> nat  
-      \<Longrightarrow> sum (%i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) =  
+      \<Longrightarrow> sum (\<lambda>i. (i $+ i $+ #1) $* (i $+ i $+ #1) $* (i $+ i $+ #1), n) =  
           $#n $* $#n $* (#2 $* $#n $* $#n $- #1)"
 by (induct_tac "n", auto)
 
 (*The sum of the first n positive integers equals n(n+1)/2.*)
 lemma sum_of_naturals:
-     "n \<in> nat \<Longrightarrow> #2 $* sum(%i. i, succ(n)) = $#n $* $#succ(n)"
+     "n \<in> nat \<Longrightarrow> #2 $* sum(\<lambda>i. i, succ(n)) = $#n $* $#succ(n)"
 by (induct_tac "n", auto)
 
 lemma sum_of_squares:
-     "n \<in> nat \<Longrightarrow> #6 $* sum (%i. i$*i, succ(n)) =  
+     "n \<in> nat \<Longrightarrow> #6 $* sum (\<lambda>i. i$*i, succ(n)) =  
                   $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1)"
 by (induct_tac "n", auto)
 
 lemma sum_of_cubes:
-     "n \<in> nat \<Longrightarrow> #4 $* sum (%i. i$*i$*i, succ(n)) =  
+     "n \<in> nat \<Longrightarrow> #4 $* sum (\<lambda>i. i$*i$*i, succ(n)) =  
                   $#n $* $#n $* ($#n $+ #1) $* ($#n $+ #1)"
 by (induct_tac "n", auto)
 
 (** Sum of fourth powers **)
 
 lemma sum_of_fourth_powers:
-     "n \<in> nat \<Longrightarrow> #30 $* sum (%i. i$*i$*i$*i, succ(n)) =  
+     "n \<in> nat \<Longrightarrow> #30 $* sum (\<lambda>i. i$*i$*i$*i, succ(n)) =  
                     $#n $* ($#n $+ #1) $* (#2 $* $#n $+ #1) $*  
                     (#3 $* $#n $* $#n $+ #3 $* $#n $- #1)"
 by (induct_tac "n", auto)
--- a/src/ZF/ex/Primes.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/Primes.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -8,22 +8,22 @@
 theory Primes imports ZF begin
 
 definition
-  divides :: "[i,i]=>o"              (infixl \<open>dvd\<close> 50)  where
+  divides :: "[i,i]\<Rightarrow>o"              (infixl \<open>dvd\<close> 50)  where
     "m dvd n \<equiv> m \<in> nat \<and> n \<in> nat \<and> (\<exists>k \<in> nat. n = m#*k)"
 
 definition
-  is_gcd  :: "[i,i,i]=>o"     \<comment> \<open>definition of great common divisor\<close>  where
+  is_gcd  :: "[i,i,i]\<Rightarrow>o"     \<comment> \<open>definition of great common divisor\<close>  where
     "is_gcd(p,m,n) \<equiv> ((p dvd m) \<and> (p dvd n))   \<and>
                        (\<forall>d\<in>nat. (d dvd m) \<and> (d dvd n) \<longrightarrow> d dvd p)"
 
 definition
-  gcd     :: "[i,i]=>i"       \<comment> \<open>Euclid's algorithm for the gcd\<close>  where
+  gcd     :: "[i,i]\<Rightarrow>i"       \<comment> \<open>Euclid's algorithm for the gcd\<close>  where
     "gcd(m,n) \<equiv> transrec(natify(n),
-                        %n f. \<lambda>m \<in> nat.
+                        \<lambda>n f. \<lambda>m \<in> nat.
                                 if n=0 then m else f`(m mod n)`n) ` natify(m)"
 
 definition
-  coprime :: "[i,i]=>o"       \<comment> \<open>the coprime relation\<close>  where
+  coprime :: "[i,i]\<Rightarrow>o"       \<comment> \<open>the coprime relation\<close>  where
     "coprime(m,n) \<equiv> gcd(m,n) = 1"
   
 definition
@@ -93,7 +93,7 @@
 lemma gcd_non_0_raw: 
     "\<lbrakk>0<n;  n \<in> nat\<rbrakk> \<Longrightarrow> gcd(m,n) = gcd(n, m mod n)"
 apply (simp add: gcd_def)
-apply (rule_tac P = "%z. left (z) = right" for left right in transrec [THEN ssubst])
+apply (rule_tac P = "\<lambda>z. left (z) = right" for left right in transrec [THEN ssubst])
 apply (simp add: ltD [THEN mem_imp_not_eq, THEN not_sym] 
                  mod_less_divisor [THEN ltD])
 done
--- a/src/ZF/ex/Ramsey.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/Ramsey.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -27,23 +27,23 @@
 theory Ramsey imports ZF begin
 
 definition
-  Symmetric :: "i=>o" where
-    "Symmetric(E) \<equiv> (\<forall>x y. <x,y>:E \<longrightarrow> <y,x>:E)"
+  Symmetric :: "i\<Rightarrow>o" where
+    "Symmetric(E) \<equiv> (\<forall>x y. \<langle>x,y\<rangle>:E \<longrightarrow> \<langle>y,x\<rangle>:E)"
 
 definition
-  Atleast :: "[i,i]=>o" where \<comment> \<open>not really necessary: ZF defines cardinality\<close>
+  Atleast :: "[i,i]\<Rightarrow>o" where \<comment> \<open>not really necessary: ZF defines cardinality\<close>
     "Atleast(n,S) \<equiv> (\<exists>f. f \<in> inj(n,S))"
 
 definition
-  Clique  :: "[i,i,i]=>o" where
-    "Clique(C,V,E) \<equiv> (C \<subseteq> V) \<and> (\<forall>x \<in> C. \<forall>y \<in> C. x\<noteq>y \<longrightarrow> <x,y> \<in> E)"
+  Clique  :: "[i,i,i]\<Rightarrow>o" where
+    "Clique(C,V,E) \<equiv> (C \<subseteq> V) \<and> (\<forall>x \<in> C. \<forall>y \<in> C. x\<noteq>y \<longrightarrow> \<langle>x,y\<rangle> \<in> E)"
 
 definition
-  Indept  :: "[i,i,i]=>o" where
-    "Indept(I,V,E) \<equiv> (I \<subseteq> V) \<and> (\<forall>x \<in> I. \<forall>y \<in> I. x\<noteq>y \<longrightarrow> <x,y> \<notin> E)"
+  Indept  :: "[i,i,i]\<Rightarrow>o" where
+    "Indept(I,V,E) \<equiv> (I \<subseteq> V) \<and> (\<forall>x \<in> I. \<forall>y \<in> I. x\<noteq>y \<longrightarrow> \<langle>x,y\<rangle> \<notin> E)"
   
 definition
-  Ramsey  :: "[i,i,i]=>o" where
+  Ramsey  :: "[i,i,i]\<Rightarrow>o" where
     "Ramsey(n,i,j) \<equiv> \<forall>V E. Symmetric(E) \<and> Atleast(n,V) \<longrightarrow>  
          (\<exists>C. Clique(C,V,E) \<and> Atleast(i,C)) |       
          (\<exists>I. Indept(I,V,E) \<and> Atleast(j,I))"
@@ -145,7 +145,7 @@
 
 (*For the Atleast part, proves \<not>(a \<in> I) from the second premise!*)
 lemma Indept_succ: 
-    "\<lbrakk>Indept(I, {z \<in> V-{a}. <a,z> \<notin> E}, E);  Symmetric(E);  a \<in> V;   
+    "\<lbrakk>Indept(I, {z \<in> V-{a}. \<langle>a,z\<rangle> \<notin> E}, E);  Symmetric(E);  a \<in> V;   
         Atleast(j,I)\<rbrakk> \<Longrightarrow>    
      Indept(cons(a,I), V, E) \<and> Atleast(succ(j), cons(a,I))"
 apply (unfold Symmetric_def Indept_def)
@@ -154,7 +154,7 @@
 
 
 lemma Clique_succ: 
-    "\<lbrakk>Clique(C, {z \<in> V-{a}. <a,z>:E}, E);  Symmetric(E);  a \<in> V;   
+    "\<lbrakk>Clique(C, {z \<in> V-{a}. \<langle>a,z\<rangle>:E}, E);  Symmetric(E);  a \<in> V;   
         Atleast(j,C)\<rbrakk> \<Longrightarrow>    
      Clique(cons(a,C), V, E) \<and> Atleast(succ(j), cons(a,C))"
 apply (unfold Symmetric_def Clique_def)
@@ -169,7 +169,7 @@
        m \<in> nat;  n \<in> nat\<rbrakk> \<Longrightarrow> Ramsey(succ(m#+n), succ(i), succ(j))"
 apply (unfold Ramsey_def, clarify)
 apply (erule Atleast_succD [THEN bexE])
-apply (erule_tac P1 = "%z.<x,z>:E" in Atleast_partition [THEN disjE],
+apply (erule_tac P1 = "\<lambda>z.\<langle>x,z\<rangle>:E" in Atleast_partition [THEN disjE],
        assumption+)
 (*case m*)
 apply (fast dest!: Indept_succ elim: Clique_superset)
--- a/src/ZF/ex/Ring.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/Ring.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -12,27 +12,27 @@
 
 (*First, we must simulate a record declaration:
 record ring = monoid +
-  add :: "[i, i] => i" (infixl "\<oplus>\<index>" 65)
+  add :: "[i, i] \<Rightarrow> i" (infixl "\<oplus>\<index>" 65)
   zero :: i ("\<zero>\<index>")
 *)
 
 definition
-  add_field :: "i => i" where
+  add_field :: "i \<Rightarrow> i" where
   "add_field(M) = fst(snd(snd(snd(M))))"
 
 definition
-  ring_add :: "[i, i, i] => i" (infixl \<open>\<oplus>\<index>\<close> 65) where
-  "ring_add(M,x,y) = add_field(M) ` <x,y>"
+  ring_add :: "[i, i, i] \<Rightarrow> i" (infixl \<open>\<oplus>\<index>\<close> 65) where
+  "ring_add(M,x,y) = add_field(M) ` \<langle>x,y\<rangle>"
 
 definition
-  zero :: "i => i" (\<open>\<zero>\<index>\<close>) where
+  zero :: "i \<Rightarrow> i" (\<open>\<zero>\<index>\<close>) where
   "zero(M) = fst(snd(snd(snd(snd(M)))))"
 
 
 lemma add_field_eq [simp]: "add_field(<C,M,I,A,z>) = A"
   by (simp add: add_field_def)
 
-lemma add_eq [simp]: "ring_add(<C,M,I,A,z>, x, y) = A ` <x,y>"
+lemma add_eq [simp]: "ring_add(<C,M,I,A,z>, x, y) = A ` \<langle>x,y\<rangle>"
   by (simp add: ring_add_def)
 
 lemma zero_eq [simp]: "zero(<C,M,I,A,Z,z>) = Z"
@@ -42,11 +42,11 @@
 text \<open>Derived operations.\<close>
 
 definition
-  a_inv :: "[i,i] => i" (\<open>\<ominus>\<index> _\<close> [81] 80) where
+  a_inv :: "[i,i] \<Rightarrow> i" (\<open>\<ominus>\<index> _\<close> [81] 80) where
   "a_inv(R) \<equiv> m_inv (<carrier(R), add_field(R), zero(R), 0>)"
 
 definition
-  minus :: "[i,i,i] => i" (\<open>(_ \<ominus>\<index> _)\<close> [65,66] 65) where
+  minus :: "[i,i,i] \<Rightarrow> i" (\<open>(_ \<ominus>\<index> _)\<close> [65,66] 65) where
   "\<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"
 
 locale abelian_monoid = fixes G (structure)
@@ -231,7 +231,7 @@
 subsection \<open>Morphisms\<close>
 
 definition
-  ring_hom :: "[i,i] => i" where
+  ring_hom :: "[i,i] \<Rightarrow> i" where
   "ring_hom(R,S) \<equiv>
     {h \<in> carrier(R) -> carrier(S).
       (\<forall>x y. x \<in> carrier(R) \<and> y \<in> carrier(R) \<longrightarrow>
--- a/src/ZF/ex/misc.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/ex/misc.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -31,7 +31,7 @@
 
 
 text\<open>A weird property of ordered pairs.\<close>
-lemma "b\<noteq>c \<Longrightarrow> <a,b> \<inter> <a,c> = <a,a>"
+lemma "b\<noteq>c \<Longrightarrow> \<langle>a,b\<rangle> \<inter> \<langle>a,c\<rangle> = \<langle>a,a\<rangle>"
 by (simp add: Pair_def Int_cons_left Int_cons_right doubleton_eq_iff, blast)
 
 text\<open>These two are cited in Benzmueller and Kohlhase's system description of
@@ -81,7 +81,7 @@
   rewriting does not instantiate Vars.*)
 lemma "(\<forall>A f B g. hom(A,f,B,g) =  
            {H \<in> A->B. f \<in> A*A->A \<and> g \<in> B*B->B \<and>  
-                     (\<forall>x \<in> A. \<forall>y \<in> A. H`(f`<x,y>) = g`<H`x,H`y>)}) \<longrightarrow>  
+                     (\<forall>x \<in> A. \<forall>y \<in> A. H`(f`\<langle>x,y\<rangle>) = g`<H`x,H`y>)}) \<longrightarrow>  
        J \<in> hom(A,f,B,g) \<and> K \<in> hom(B,g,C,h) \<longrightarrow>   
        (K O J) \<in> hom(A,f,C,h)"
 by force
@@ -89,7 +89,7 @@
 text\<open>Another version, with meta-level rewriting\<close>
 lemma "(\<And>A f B g. hom(A,f,B,g) \<equiv>  
            {H \<in> A->B. f \<in> A*A->A \<and> g \<in> B*B->B \<and>  
-                     (\<forall>x \<in> A. \<forall>y \<in> A. H`(f`<x,y>) = g`<H`x,H`y>)}) 
+                     (\<forall>x \<in> A. \<forall>y \<in> A. H`(f`\<langle>x,y\<rangle>) = g`<H`x,H`y>)}) 
        \<Longrightarrow> J \<in> hom(A,f,B,g) \<and> K \<in> hom(B,g,C,h) \<longrightarrow> (K O J) \<in> hom(A,f,C,h)"
 by force
 
--- a/src/ZF/func.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/func.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -25,7 +25,7 @@
 
 (*For upward compatibility with the former definition*)
 lemma Pi_iff_old:
-    "f \<in> Pi(A,B) \<longleftrightarrow> f<=Sigma(A,B) \<and> (\<forall>x\<in>A. \<exists>!y. <x,y>: f)"
+    "f \<in> Pi(A,B) \<longleftrightarrow> f<=Sigma(A,B) \<and> (\<forall>x\<in>A. \<exists>!y. \<langle>x,y\<rangle>: f)"
 by (unfold Pi_def function_def, blast)
 
 lemma fun_is_function: "f \<in> Pi(A,B) \<Longrightarrow> function(f)"
@@ -36,7 +36,7 @@
 by (simp add: Pi_iff relation_def, blast)
 
 lemma functionI:
-     "\<lbrakk>\<And>x y y'. \<lbrakk><x,y>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
+     "\<lbrakk>\<And>x y y'. \<lbrakk>\<langle>x,y\<rangle>:r; <x,y'>:r\<rbrakk> \<Longrightarrow> y=y'\<rbrakk> \<Longrightarrow> function(r)"
 by (simp add: function_def, blast)
 
 (*Functions are relations*)
@@ -57,13 +57,13 @@
 
 subsection\<open>Function Application\<close>
 
-lemma apply_equality2: "\<lbrakk><a,b>: f;  <a,c>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b=c"
+lemma apply_equality2: "\<lbrakk>\<langle>a,b\<rangle>: f;  \<langle>a,c\<rangle>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b=c"
 by (unfold Pi_def function_def, blast)
 
-lemma function_apply_equality: "\<lbrakk><a,b>: f;  function(f)\<rbrakk> \<Longrightarrow> f`a = b"
+lemma function_apply_equality: "\<lbrakk>\<langle>a,b\<rangle>: f;  function(f)\<rbrakk> \<Longrightarrow> f`a = b"
 by (unfold apply_def function_def, blast)
 
-lemma apply_equality: "\<lbrakk><a,b>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> f`a = b"
+lemma apply_equality: "\<lbrakk>\<langle>a,b\<rangle>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> f`a = b"
 apply (unfold Pi_def)
 apply (blast intro: function_apply_equality)
 done
@@ -96,7 +96,7 @@
 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) \<Longrightarrow> <a,b>: f \<longleftrightarrow> a \<in> A \<and> f`a = b"
+lemma apply_iff: "f \<in> Pi(A,B) \<Longrightarrow> \<langle>a,b\<rangle>: f \<longleftrightarrow> a \<in> A \<and> f`a = b"
 apply (frule fun_is_rel)
 apply (blast intro!: apply_Pair apply_equality)
 done
@@ -109,7 +109,7 @@
 
 (*Such functions arise in non-standard datatypes, ZF/ex/Ntree for instance*)
 lemma Pi_Collect_iff:
-     "(f \<in> Pi(A, %x. {y \<in> B(x). P(x,y)}))
+     "(f \<in> Pi(A, \<lambda>x. {y \<in> B(x). P(x,y)}))
       \<longleftrightarrow>  f \<in> Pi(A,B) \<and> (\<forall>x\<in>A. P(x, f`x))"
 by (blast intro: Pi_type dest: apply_type)
 
@@ -120,13 +120,13 @@
 
 (** Elimination of membership in a function **)
 
-lemma domain_type: "\<lbrakk><a,b> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A"
+lemma domain_type: "\<lbrakk>\<langle>a,b\<rangle> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A"
 by (blast dest: fun_is_rel)
 
-lemma range_type: "\<lbrakk><a,b> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b \<in> B(a)"
+lemma range_type: "\<lbrakk>\<langle>a,b\<rangle> \<in> f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> b \<in> B(a)"
 by (blast dest: fun_is_rel)
 
-lemma Pair_mem_PiD: "\<lbrakk><a,b>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> B(a) \<and> f`a = b"
+lemma Pair_mem_PiD: "\<lbrakk>\<langle>a,b\<rangle>: f;  f \<in> Pi(A,B)\<rbrakk> \<Longrightarrow> a \<in> A \<and> b \<in> B(a) \<and> f`a = b"
 by (blast intro: domain_type range_type apply_equality)
 
 subsection\<open>Lambda Abstraction\<close>
@@ -141,7 +141,7 @@
 \<rbrakk> \<Longrightarrow>  P"
 by (simp add: lam_def, blast)
 
-lemma lamD: "\<lbrakk><a,c>: (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> c = b(a)"
+lemma lamD: "\<lbrakk>\<langle>a,c\<rangle>: (\<lambda>x\<in>A. b(x))\<rbrakk> \<Longrightarrow> c = b(a)"
 by (simp add: lam_def)
 
 lemma lam_type [TC]:
@@ -190,7 +190,7 @@
 by (unfold Pi_def function_def, blast)
 
 (*The singleton function*)
-lemma singleton_fun [simp]: "{<a,b>} \<in> {a} -> {b}"
+lemma singleton_fun [simp]: "{\<langle>a,b\<rangle>} \<in> {a} -> {b}"
 by (unfold Pi_def function_def, blast)
 
 lemma Pi_empty2 [simp]: "(A->0) = (if A=0 then {0} else 0)"
@@ -395,21 +395,21 @@
 subsection\<open>Extensions of Functions\<close>
 
 lemma fun_extend:
-     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(<c,b>,f) \<in> cons(c,A) -> cons(b,B)"
+     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(\<langle>c,b\<rangle>,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:
-     "\<lbrakk>f \<in> A->B;  c\<notin>A;  b \<in> B\<rbrakk> \<Longrightarrow> cons(<c,b>,f) \<in> cons(c,A) -> B"
+     "\<lbrakk>f \<in> A->B;  c\<notin>A;  b \<in> B\<rbrakk> \<Longrightarrow> cons(\<langle>c,b\<rangle>,f) \<in> cons(c,A) -> B"
 by (blast intro: fun_extend [THEN fun_weaken_type])
 
 lemma extend_apply:
-     "c \<notin> domain(f) \<Longrightarrow> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+     "c \<notin> domain(f) \<Longrightarrow> cons(\<langle>c,b\<rangle>,f)`a = (if a=c then b else f`a)"
 by (auto simp add: apply_def)
 
 lemma fun_extend_apply [simp]:
-     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(<c,b>,f)`a = (if a=c then b else f`a)"
+     "\<lbrakk>f \<in> A->B;  c\<notin>A\<rbrakk> \<Longrightarrow> cons(\<langle>c,b\<rangle>,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 \<Longrightarrow> 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(\<langle>c,b\<rangle>, f)})"
 apply (rule equalityI)
 apply (safe elim!: fun_extend3)
 (*Inclusion of left into right*)
@@ -435,14 +435,14 @@
 apply (erule consE, simp_all)
 done
 
-lemma succ_fun_eq: "succ(n) -> B = (\<Union>f \<in> n->B. \<Union>b\<in>B. {cons(<n,b>, f)})"
+lemma succ_fun_eq: "succ(n) -> B = (\<Union>f \<in> n->B. \<Union>b\<in>B. {cons(\<langle>n,b\<rangle>, f)})"
 by (simp add: succ_def mem_not_refl cons_fun_eq)
 
 
 subsection\<open>Function Updates\<close>
 
 definition
-  update  :: "[i,i,i] => i"  where
+  update  :: "[i,i,i] \<Rightarrow> i"  where
    "update(f,a,b) \<equiv> \<lambda>x\<in>cons(a, domain(f)). if(x=a, b, f`x)"
 
 nonterminal updbinds and updbind
@@ -451,10 +451,10 @@
 
   (* Let expressions *)
 
-  "_updbind"    :: "[i, i] => updbind"               (\<open>(2_ :=/ _)\<close>)
-  ""            :: "updbind => updbinds"             (\<open>_\<close>)
-  "_updbinds"   :: "[updbind, updbinds] => updbinds" (\<open>_,/ _\<close>)
-  "_Update"     :: "[i, updbinds] => i"              (\<open>_/'((_)')\<close> [900,0] 900)
+  "_updbind"    :: "[i, i] \<Rightarrow> updbind"               (\<open>(2_ :=/ _)\<close>)
+  ""            :: "updbind \<Rightarrow> updbinds"             (\<open>_\<close>)
+  "_updbinds"   :: "[updbind, updbinds] \<Rightarrow> updbinds" (\<open>_,/ _\<close>)
+  "_Update"     :: "[i, updbinds] \<Rightarrow> i"              (\<open>_/'((_)')\<close> [900,0] 900)
 
 translations
   "_Update (f, _updbinds(b,bs))"  == "_Update (_Update(f,b), bs)"
@@ -568,11 +568,11 @@
 subsubsection\<open>Images\<close>
 
 lemma image_pair_mono:
-    "\<lbrakk>\<And>x y. <x,y>:r \<Longrightarrow> <x,y>:s;  A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
+    "\<lbrakk>\<And>x y. \<langle>x,y\<rangle>:r \<Longrightarrow> \<langle>x,y\<rangle>:s;  A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
 by blast
 
 lemma vimage_pair_mono:
-    "\<lbrakk>\<And>x y. <x,y>:r \<Longrightarrow> <x,y>:s;  A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
+    "\<lbrakk>\<And>x y. \<langle>x,y\<rangle>:r \<Longrightarrow> \<langle>x,y\<rangle>:s;  A<=B\<rbrakk> \<Longrightarrow> r-``A \<subseteq> s-``B"
 by blast
 
 lemma image_mono: "\<lbrakk>r<=s;  A<=B\<rbrakk> \<Longrightarrow> r``A \<subseteq> s``B"
--- a/src/ZF/pair.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/pair.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -29,7 +29,7 @@
 \<close>
 
 
-(** Lemmas for showing that <a,b> uniquely determines a and b **)
+(** Lemmas for showing that \<langle>a,b\<rangle> uniquely determines a and b **)
 
 lemma singleton_eq_iff [iff]: "{a} = {b} \<longleftrightarrow> a=b"
 by (rule extension [THEN iff_trans], blast)
@@ -37,7 +37,7 @@
 lemma doubleton_eq_iff: "{a,b} = {c,d} \<longleftrightarrow> (a=c \<and> b=d) | (a=d \<and> b=c)"
 by (rule extension [THEN iff_trans], blast)
 
-lemma Pair_iff [simp]: "<a,b> = <c,d> \<longleftrightarrow> a=c \<and> b=d"
+lemma Pair_iff [simp]: "\<langle>a,b\<rangle> = \<langle>c,d\<rangle> \<longleftrightarrow> a=c \<and> b=d"
 by (simp add: Pair_def doubleton_eq_iff, blast)
 
 lemmas Pair_inject = Pair_iff [THEN iffD1, THEN conjE, elim!]
@@ -45,7 +45,7 @@
 lemmas Pair_inject1 = Pair_iff [THEN iffD1, THEN conjunct1]
 lemmas Pair_inject2 = Pair_iff [THEN iffD1, THEN conjunct2]
 
-lemma Pair_not_0: "<a,b> \<noteq> 0"
+lemma Pair_not_0: "\<langle>a,b\<rangle> \<noteq> 0"
 apply (unfold Pair_def)
 apply (blast elim: equalityE)
 done
@@ -54,7 +54,7 @@
 
 declare sym [THEN Pair_neq_0, elim!]
 
-lemma Pair_neq_fst: "<a,b>=a \<Longrightarrow> P"
+lemma Pair_neq_fst: "\<langle>a,b\<rangle>=a \<Longrightarrow> P"
 proof (unfold Pair_def)
   assume eq: "{{a, a}, {a, b}} = a"
   have  "{a, a} \<in> {{a, a}, {a, b}}" by (rule consI1)
@@ -63,7 +63,7 @@
   ultimately show "P" by (rule mem_asym)
 qed
 
-lemma Pair_neq_snd: "<a,b>=b \<Longrightarrow> P"
+lemma Pair_neq_snd: "\<langle>a,b\<rangle>=b \<Longrightarrow> P"
 proof (unfold Pair_def)
   assume eq: "{{a, a}, {a, b}} = b"
   have  "{a, b} \<in> {{a, a}, {a, b}}" by blast
@@ -77,10 +77,10 @@
 
 text\<open>Generalizes Cartesian product\<close>
 
-lemma Sigma_iff [simp]: "<a,b>: Sigma(A,B) \<longleftrightarrow> a \<in> A \<and> b \<in> B(a)"
+lemma Sigma_iff [simp]: "\<langle>a,b\<rangle>: Sigma(A,B) \<longleftrightarrow> a \<in> A \<and> b \<in> B(a)"
 by (simp add: Sigma_def)
 
-lemma SigmaI [TC,intro!]: "\<lbrakk>a \<in> A;  b \<in> B(a)\<rbrakk> \<Longrightarrow> <a,b> \<in> Sigma(A,B)"
+lemma SigmaI [TC,intro!]: "\<lbrakk>a \<in> A;  b \<in> B(a)\<rbrakk> \<Longrightarrow> \<langle>a,b\<rangle> \<in> Sigma(A,B)"
 by simp
 
 lemmas SigmaD1 = Sigma_iff [THEN iffD1, THEN conjunct1]
@@ -89,12 +89,12 @@
 (*The general elimination rule*)
 lemma SigmaE [elim!]:
     "\<lbrakk>c \<in> Sigma(A,B);
-        \<And>x y.\<lbrakk>x \<in> A;  y \<in> B(x);  c=<x,y>\<rbrakk> \<Longrightarrow> P
+        \<And>x y.\<lbrakk>x \<in> A;  y \<in> B(x);  c=\<langle>x,y\<rangle>\<rbrakk> \<Longrightarrow> P
 \<rbrakk> \<Longrightarrow> P"
 by (unfold Sigma_def, blast)
 
 lemma SigmaE2 [elim!]:
-    "\<lbrakk><a,b> \<in> Sigma(A,B);
+    "\<lbrakk>\<langle>a,b\<rangle> \<in> Sigma(A,B);
         \<lbrakk>a \<in> A;  b \<in> B(a)\<rbrakk> \<Longrightarrow> P
 \<rbrakk> \<Longrightarrow> P"
 by (unfold Sigma_def, blast)
@@ -120,10 +120,10 @@
 
 subsection\<open>Projections \<^term>\<open>fst\<close> and \<^term>\<open>snd\<close>\<close>
 
-lemma fst_conv [simp]: "fst(<a,b>) = a"
+lemma fst_conv [simp]: "fst(\<langle>a,b\<rangle>) = a"
 by (simp add: fst_def)
 
-lemma snd_conv [simp]: "snd(<a,b>) = b"
+lemma snd_conv [simp]: "snd(\<langle>a,b\<rangle>) = b"
 by (simp add: snd_def)
 
 lemma fst_type [TC]: "p \<in> Sigma(A,B) \<Longrightarrow> fst(p) \<in> A"
@@ -139,33 +139,33 @@
 subsection\<open>The Eliminator, \<^term>\<open>split\<close>\<close>
 
 (*A META-equality, so that it applies to higher types as well...*)
-lemma split [simp]: "split(%x y. c(x,y), <a,b>) \<equiv> c(a,b)"
+lemma split [simp]: "split(\<lambda>x y. c(x,y), \<langle>a,b\<rangle>) \<equiv> c(a,b)"
 by (simp add: split_def)
 
 lemma split_type [TC]:
     "\<lbrakk>p \<in> Sigma(A,B);
-         \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(<x,y>)
-\<rbrakk> \<Longrightarrow> split(%x y. c(x,y), p) \<in> C(p)"
+         \<And>x y.\<lbrakk>x \<in> A; y \<in> B(x)\<rbrakk> \<Longrightarrow> c(x,y):C(\<langle>x,y\<rangle>)
+\<rbrakk> \<Longrightarrow> split(\<lambda>x y. c(x,y), p) \<in> C(p)"
 by (erule SigmaE, auto)
 
 lemma expand_split:
   "u \<in> A*B \<Longrightarrow>
-        R(split(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = <x,y> \<longrightarrow> R(c(x,y)))"
+        R(split(c,u)) \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>B. u = \<langle>x,y\<rangle> \<longrightarrow> R(c(x,y)))"
 by (auto simp add: split_def)
 
 
 subsection\<open>A version of \<^term>\<open>split\<close> for Formulae: Result Type \<^typ>\<open>o\<close>\<close>
 
-lemma splitI: "R(a,b) \<Longrightarrow> split(R, <a,b>)"
+lemma splitI: "R(a,b) \<Longrightarrow> split(R, \<langle>a,b\<rangle>)"
 by (simp add: split_def)
 
 lemma splitE:
     "\<lbrakk>split(R,z);  z \<in> Sigma(A,B);
-        \<And>x y. \<lbrakk>z = <x,y>;  R(x,y)\<rbrakk> \<Longrightarrow> P
+        \<And>x y. \<lbrakk>z = \<langle>x,y\<rangle>;  R(x,y)\<rbrakk> \<Longrightarrow> P
 \<rbrakk> \<Longrightarrow> P"
 by (auto simp add: split_def)
 
-lemma splitD: "split(R,<a,b>) \<Longrightarrow> R(a,b)"
+lemma splitD: "split(R,\<langle>a,b\<rangle>) \<Longrightarrow> R(a,b)"
 by (simp add: split_def)
 
 text \<open>
@@ -173,11 +173,11 @@
 \<close>
 
 lemma split_paired_Bex_Sigma [simp]:
-     "(\<exists>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<exists>x \<in> A. \<exists>y \<in> B(x). P(<x,y>))"
+     "(\<exists>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<exists>x \<in> A. \<exists>y \<in> B(x). P(\<langle>x,y\<rangle>))"
 by blast
 
 lemma split_paired_Ball_Sigma [simp]:
-     "(\<forall>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<forall>x \<in> A. \<forall>y \<in> B(x). P(<x,y>))"
+     "(\<forall>z \<in> Sigma(A,B). P(z)) \<longleftrightarrow> (\<forall>x \<in> A. \<forall>y \<in> B(x). P(\<langle>x,y\<rangle>))"
 by blast
 
 end
--- a/src/ZF/upair.thy	Tue Sep 27 17:03:23 2022 +0100
+++ b/src/ZF/upair.thy	Tue Sep 27 17:46:52 2022 +0100
@@ -229,11 +229,11 @@
 (** Rewrite rules for boolean case-splitting: faster than split_if [split]
 **)
 
-lemmas split_if_eq1 = split_if [of "%x. x = b"] for b
-lemmas split_if_eq2 = split_if [of "%x. a = x"] for a
+lemmas split_if_eq1 = split_if [of "\<lambda>x. x = b"] for b
+lemmas split_if_eq2 = split_if [of "\<lambda>x. a = x"] for a
 
-lemmas split_if_mem1 = split_if [of "%x. x \<in> b"] for b
-lemmas split_if_mem2 = split_if [of "%x. a \<in> x"] for a
+lemmas split_if_mem1 = split_if [of "\<lambda>x. x \<in> b"] for b
+lemmas split_if_mem2 = split_if [of "\<lambda>x. a \<in> x"] for a
 
 lemmas split_ifs = split_if_eq1 split_if_eq2 split_if_mem1 split_if_mem2